diff --git a/doc/dev/ffi.md b/doc/dev/ffi.md index 29a3c14ba687..934f7ff2b5ca 100644 --- a/doc/dev/ffi.md +++ b/doc/dev/ffi.md @@ -34,7 +34,8 @@ In the case of `@[extern]` all *irrelevant* types are removed first; see next se ### Translating Types from Lean to C -* The integer types `UInt8`, ..., `UInt64`, `USize` are represented by the C types `uint8_t`, ..., `uint64_t`, `size_t`, respectively +* The integer types `UInt8`, ..., `UInt64`, `USize`, `Int8`, ..., `Int64`, `ISize` are represented by + the C types `uint8_t`, ..., `uint64_t`, `size_t`, `int8_t`, ..., `int64_t`, `ptrdiff_t` respectively * `Char` is represented by `uint32_t` * `Float` is represented by `double` * An *enum* inductive type of at least 2 and at most 2^32 constructors, each of which with no parameters, is represented by the first type of `uint8_t`, `uint16_t`, `uint32_t` that is sufficient to represent all constructor indices. @@ -49,7 +50,6 @@ In the case of `@[extern]` all *irrelevant* types are removed first; see next se is represented by the representation of that parameter's type. For example, `{ x : α // p }`, the `Subtype` structure of a value of type `α` and an irrelevant proof, is represented by the representation of `α`. - Similarly, the signed integer types `Int8`, ..., `Int64`, `ISize` are also represented by the unsigned C types `uint8_t`, ..., `uint64_t`, `size_t`, respectively, because they have a trivial structure. * `Nat` and `Int` are represented by `lean_object *`. Their runtime values is either a pointer to an opaque bignum object or, if the lowest bit of the "pointer" is 1 (`lean_is_scalar`), an encoded unboxed natural number or integer (`lean_box`/`lean_unbox`). * A universe `Sort u`, type constructor `... → Sort u`, `Void α` or proposition `p : Prop` is *irrelevant* and is either statically erased (see above) or represented as a `lean_object *` with the runtime value `lean_box(0)` diff --git a/src/Init/Data/SInt/Basic.lean b/src/Init/Data/SInt/Basic.lean index f60579ced8b3..178be26d72a2 100644 --- a/src/Init/Data/SInt/Basic.lean +++ b/src/Init/Data/SInt/Basic.lean @@ -30,6 +30,9 @@ structure Int8 where -/ toUInt8 : UInt8 +attribute [extern "lean_int8_of_uint8_mk"] Int8.ofUInt8 +attribute [extern "lean_int8_to_uint8"] Int8.toUInt8 + /-- Signed 16-bit integers. @@ -43,6 +46,9 @@ structure Int16 where -/ toUInt16 : UInt16 +attribute [extern "lean_int16_of_uint16_mk"] Int16.ofUInt16 +attribute [extern "lean_int16_to_uint16"] Int16.toUInt16 + /-- Signed 32-bit integers. @@ -56,6 +62,9 @@ structure Int32 where -/ toUInt32 : UInt32 +attribute [extern "lean_int32_of_uint32_mk"] Int32.ofUInt32 +attribute [extern "lean_int32_to_uint32"] Int32.toUInt32 + /-- Signed 64-bit integers. @@ -69,6 +78,9 @@ structure Int64 where -/ toUInt64 : UInt64 +attribute [extern "lean_int64_of_uint64_mk"] Int64.ofUInt64 +attribute [extern "lean_int64_to_uint64"] Int64.toUInt64 + /-- Signed integers that are the size of a word on the platform's architecture. @@ -83,6 +95,9 @@ structure ISize where -/ toUSize : USize +attribute [extern "lean_isize_of_usize_mk"] ISize.ofUSize +attribute [extern "lean_isize_to_usize"] ISize.toUSize + /-- The number of distinct values representable by `Int8`, that is, `2^8 = 256`. -/ abbrev Int8.size : Nat := 256 diff --git a/src/Lean/Compiler/IR/Basic.lean b/src/Lean/Compiler/IR/Basic.lean index 3deace0e7709..09b95c2b84d1 100644 --- a/src/Lean/Compiler/IR/Basic.lean +++ b/src/Lean/Compiler/IR/Basic.lean @@ -82,6 +82,7 @@ inductive IRType where -- TODO: Move this upwards after a stage0 update. | tagged | void + | int8 | int16 | int32 | int64 | isize deriving Inhabited, BEq, Repr namespace IRType @@ -94,6 +95,11 @@ def isScalar : IRType → Bool | uint32 => true | uint64 => true | usize => true + | int8 => true + | int16 => true + | int32 => true + | int64 => true + | isize => true | _ => false def isObj : IRType → Bool @@ -124,6 +130,9 @@ def boxed : IRType → IRType | void | tagged | uint8 | uint16 => tagged | _ => tobject +def ptrSizedTypeForSign (signed : Bool) : IRType := + if signed then .isize else .usize + end IRType /-- Arguments to applications, constructors, etc. @@ -185,7 +194,7 @@ inductive Expr where 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) + | uproj (i : Nat) (signed : Bool) (x : VarId) /-- Extract the scalar value at Position `sizeof(void*)*n + offset` from `x`. -/ | sproj (n : Nat) (offset : Nat) (x : VarId) /-- Full application. -/ @@ -226,7 +235,7 @@ inductive FnBody where | 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) + | uset (x : VarId) (i : Nat) (y : VarId) (signed : Bool) (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) @@ -261,7 +270,7 @@ def FnBody.body : FnBody → FnBody | FnBody.vdecl _ _ _ b => b | FnBody.jdecl _ _ _ b => b | FnBody.set _ _ _ b => b - | FnBody.uset _ _ _ b => b + | FnBody.uset _ _ _ _ b => b | FnBody.sset _ _ _ _ _ b => b | FnBody.setTag _ _ b => b | FnBody.inc _ _ _ _ b => b @@ -273,7 +282,7 @@ 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.uset x i y s _, b => FnBody.uset x i y s 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 @@ -484,7 +493,7 @@ def Expr.alphaEqv (ρ : IndexRenaming) : Expr → Expr → Bool | 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.uproj i₁ s₁ x₁, Expr.uproj i₂ s₂ x₂ => i₁ == i₂ && s₁ == s₂ && 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₂ @@ -521,7 +530,7 @@ partial def FnBody.alphaEqv : IndexRenaming → FnBody → FnBody → Bool | 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.uset x₁ i₁ y₁ s₁ b₁, FnBody.uset x₂ i₂ y₂ s₂ b₂ => aeqv ρ x₁ x₂ && i₁ == i₂ && aeqv ρ y₁ y₂ && s₁ == s₂ && 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₂ @@ -556,6 +565,9 @@ def getUnboxOpName (t : IRType) : String := | IRType.usize => "lean_unbox_usize" | IRType.uint32 => "lean_unbox_uint32" | IRType.uint64 => "lean_unbox_uint64" + | IRType.isize => "lean_unbox_isize" + | IRType.int32 => "lean_unbox_int32" + | IRType.int64 => "lean_unbox_int64" | IRType.float => "lean_unbox_float" | IRType.float32 => "lean_unbox_float32" | _ => "lean_unbox" diff --git a/src/Lean/Compiler/IR/Boxing.lean b/src/Lean/Compiler/IR/Boxing.lean index e2867293badc..ca819c53bff7 100644 --- a/src/Lean/Compiler/IR/Boxing.lean +++ b/src/Lean/Compiler/IR/Boxing.lean @@ -281,10 +281,10 @@ 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 i y s b => do let b ← visitFnBody b - castVarIfNeeded y IRType.usize fun y => - return .uset x i y b + castVarIfNeeded y (IRType.ptrSizedTypeForSign s) fun y => + return .uset x i y s b | .sset x i o y ty b => do let b ← visitFnBody b castVarIfNeeded y ty fun y => diff --git a/src/Lean/Compiler/IR/Checker.lean b/src/Lean/Compiler/IR/Checker.lean index 1d391ada3091..cc0d691de730 100644 --- a/src/Lean/Compiler/IR/Checker.lean +++ b/src/Lean/Compiler/IR/Checker.lean @@ -168,9 +168,9 @@ def checkExpr (ty : IRType) (e : Expr) : M Unit := do else throwCheckerError "invalid proj index" | _ => throwCheckerError s!"unexpected IR type '{xType}'" - | .uproj _ x => + | .uproj _ s x => checkObjVar x - checkType ty (· == .usize) + checkType ty (· == IRType.ptrSizedTypeForSign s) | .sproj _ _ x => checkObjVar x checkScalarType ty @@ -202,7 +202,7 @@ 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 diff --git a/src/Lean/Compiler/IR/EmitC.lean b/src/Lean/Compiler/IR/EmitC.lean index 03bc22778daf..e783aede8f3c 100644 --- a/src/Lean/Compiler/IR/EmitC.lean +++ b/src/Lean/Compiler/IR/EmitC.lean @@ -61,6 +61,11 @@ def toCType : IRType → String | IRType.uint32 => "uint32_t" | IRType.uint64 => "uint64_t" | IRType.usize => "size_t" + | IRType.int8 => "int8_t" + | IRType.int16 => "int16_t" + | IRType.int32 => "int32_t" + | IRType.int64 => "int64_t" + | IRType.isize => "ptrdiff_t" | IRType.object => "lean_object*" | IRType.tagged => "lean_object*" | IRType.tobject => "lean_object*" @@ -314,8 +319,13 @@ 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 emitUSet (x : VarId) (n : Nat) (s : Bool) (y : VarId) : M Unit := do + if s then + emit "lean_ctor_set_isize("; + else + 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 match t with @@ -325,6 +335,10 @@ def emitSSet (x : VarId) (n : Nat) (offset : Nat) (y : VarId) (t : IRType) : M U | IRType.uint16 => emit "lean_ctor_set_uint16" | IRType.uint32 => emit "lean_ctor_set_uint32" | IRType.uint64 => emit "lean_ctor_set_uint64" + | IRType.int8 => emit "lean_ctor_set_int8" + | IRType.int16 => emit "lean_ctor_set_int16" + | IRType.int32 => emit "lean_ctor_set_int32" + | IRType.int64 => emit "lean_ctor_set_int64" | _ => throw "invalid instruction"; emit "("; emit x; emit ", "; emitOffset n offset; emit ", "; emit y; emitLn ");" @@ -389,8 +403,13 @@ def emitReuse (z : VarId) (x : VarId) (c : CtorInfo) (updtHeader : Bool) (ys : A def emitProj (z : VarId) (i : Nat) (x : VarId) : M Unit := do emitLhs z; 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) (i : Nat) (s : Bool) (x : VarId) : M Unit := do + emitLhs z; + if s then + emit "lean_ctor_get_isize(" + 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; @@ -401,6 +420,10 @@ def emitSProj (z : VarId) (t : IRType) (n offset : Nat) (x : VarId) : M Unit := | IRType.uint16 => emit "lean_ctor_get_uint16" | IRType.uint32 => emit "lean_ctor_get_uint32" | IRType.uint64 => emit "lean_ctor_get_uint64" + | IRType.int8 => emit "lean_ctor_get_int8" + | IRType.int16 => emit "lean_ctor_get_int16" + | IRType.int32 => emit "lean_ctor_get_int32" + | IRType.int64 => emit "lean_ctor_get_int64" | _ => throw "invalid instruction" emit "("; emit x; emit ", "; emitOffset n offset; emitLn ");" @@ -461,6 +484,9 @@ def emitBoxFn (xType : IRType) : M Unit := | IRType.usize => emit "lean_box_usize" | IRType.uint32 => emit "lean_box_uint32" | IRType.uint64 => emit "lean_box_uint64" + | IRType.isize => emit "lean_box_isize" + | IRType.int32 => emit "lean_box_int32" + | IRType.int64 => emit "lean_box_int64" | IRType.float => emit "lean_box_float" | IRType.float32 => emit "lean_box_float32" | _ => emit "lean_box" @@ -529,7 +555,7 @@ def emitVDecl (z : VarId) (t : IRType) (v : Expr) : M Unit := | 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.uproj i s x => emitUProj z i s x | Expr.sproj n o x => emitSProj z t n o x | Expr.fap c ys => emitFullApp z c ys | Expr.pap c ys => emitPartialApp z c ys @@ -639,7 +665,7 @@ 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.uset x i y s b => emitUSet x i s y; emitBlock b | FnBody.sset x i o y t b => emitSSet x i o y t; emitBlock b | FnBody.ret x => emit "return "; emitArg x; emitLn ";" | FnBody.case _ x xType alts => emitCase x xType alts diff --git a/src/Lean/Compiler/IR/EmitLLVM.lean b/src/Lean/Compiler/IR/EmitLLVM.lean index 0a30d8587efa..040e96c417fa 100644 --- a/src/Lean/Compiler/IR/EmitLLVM.lean +++ b/src/Lean/Compiler/IR/EmitLLVM.lean @@ -30,6 +30,8 @@ time. These changes can likely be done similar to the ones in EmitC: - function decls need to be fixed - full applications need to be fixed - tail calls need to be fixed +- signed integer introduction: + - mirror changes for boxing/unboxing IntX from EmitC -/ def leanMainFn := "_lean_main" @@ -333,6 +335,12 @@ def toLLVMType (t : IRType) : M llvmctx (LLVM.LLVMType llvmctx) := do | IRType.uint64 => LLVM.intTypeInContext llvmctx 64 -- TODO: how to cleanly size_t in LLVM? We can do eg. instantiate the current target and query for size. | IRType.usize => LLVM.size_tType llvmctx + | IRType.int8 => LLVM.intTypeInContext llvmctx 8 + | IRType.int16 => LLVM.intTypeInContext llvmctx 16 + | IRType.int32 => LLVM.intTypeInContext llvmctx 32 + | IRType.int64 => LLVM.intTypeInContext llvmctx 64 + -- TODO: how to cleanly size_t in LLVM? We can do eg. instantiate the current target and query for size. + | IRType.isize => LLVM.size_tType llvmctx | IRType.object => do LLVM.pointerType (← LLVM.i8Type llvmctx) | IRType.tagged => do LLVM.pointerType (← LLVM.i8Type llvmctx) | IRType.tobject => do LLVM.pointerType (← LLVM.i8Type llvmctx) @@ -806,18 +814,19 @@ def emitProj (builder : LLVM.Builder llvmctx) (z : VarId) (i : Nat) (x : VarId) let zval ← callLeanCtorGet builder xval (← constIntUnsigned i) "" emitLhsSlotStore builder z zval -def callLeanCtorGetUsize (builder : LLVM.Builder llvmctx) - (x i : LLVM.Value llvmctx) (retName : String) : M llvmctx (LLVM.Value llvmctx) := do - let fnName := "lean_ctor_get_usize" +def callLeanCtorGetPtrSizedType (builder : LLVM.Builder llvmctx) + (x i : LLVM.Value llvmctx) (s : Bool) (retName : String) : M llvmctx (LLVM.Value llvmctx) := do + let fnName := if s then "lean_ctor_get_isize" else "lean_ctor_get_usize" let retty ← LLVM.size_tType llvmctx let argtys := #[ ← LLVM.voidPtrType llvmctx, ← LLVM.unsignedType llvmctx] let fnty ← LLVM.functionType retty argtys let fn ← getOrCreateFunctionPrototype (← getLLVMModule) retty fnName argtys LLVM.buildCall2 builder fnty fn #[x, i] retName -def emitUProj (builder : LLVM.Builder llvmctx) (z : VarId) (i : Nat) (x : VarId) : M llvmctx Unit := do +def emitUProj (builder : LLVM.Builder llvmctx) (z : VarId) (i : Nat) (s : Bool) + (x : VarId) : M llvmctx Unit := do let xval ← emitLhsVal builder x - let zval ← callLeanCtorGetUsize builder xval (← constIntUnsigned i) "" + let zval ← callLeanCtorGetPtrSizedType builder xval (← constIntUnsigned i) s "" emitLhsSlotStore builder z zval def emitOffset (builder : LLVM.Builder llvmctx) @@ -837,6 +846,10 @@ def emitSProj (builder : LLVM.Builder llvmctx) | IRType.uint16 => pure ("lean_ctor_get_uint16", ← LLVM.i16Type llvmctx) | IRType.uint32 => pure ("lean_ctor_get_uint32", ← LLVM.i32Type llvmctx) | IRType.uint64 => pure ("lean_ctor_get_uint64", ← LLVM.i64Type llvmctx) + | IRType.int8 => pure ("lean_ctor_get_int8", ← LLVM.i8Type llvmctx) + | IRType.int16 => pure ("lean_ctor_get_int16", ← LLVM.i16Type llvmctx) + | IRType.int32 => pure ("lean_ctor_get_int32", ← LLVM.i32Type llvmctx) + | IRType.int64 => pure ("lean_ctor_get_int64", ← LLVM.i64Type llvmctx) | _ => throw s!"Invalid type for lean_ctor_get: '{t}'" let argtys := #[ ← LLVM.voidPtrType llvmctx, ← LLVM.unsignedType llvmctx] let fn ← getOrCreateFunctionPrototype (← getLLVMModule) retty fnName argtys @@ -880,6 +893,9 @@ def emitBox (builder : LLVM.Builder llvmctx) (z : VarId) (x : VarId) (xType : IR | IRType.usize => pure ("lean_box_usize", ← LLVM.size_tType llvmctx, xv) | IRType.uint32 => pure ("lean_box_uint32", ← LLVM.i32Type llvmctx, xv) | IRType.uint64 => pure ("lean_box_uint64", ← LLVM.size_tType llvmctx, xv) + | IRType.isize => pure ("lean_box_isize", ← LLVM.size_tType llvmctx, xv) + | IRType.int32 => pure ("lean_box_int32", ← LLVM.i32Type llvmctx, xv) + | IRType.int64 => pure ("lean_box_int64", ← LLVM.size_tType llvmctx, xv) | IRType.float => pure ("lean_box_float", ← LLVM.doubleTypeInContext llvmctx, xv) | IRType.float32 => pure ("lean_box_float32", ← LLVM.floatTypeInContext llvmctx, xv) | _ => @@ -911,6 +927,9 @@ def callUnboxForType (builder : LLVM.Builder llvmctx) | IRType.usize => pure ("lean_unbox_usize", ← toLLVMType t) | IRType.uint32 => pure ("lean_unbox_uint32", ← toLLVMType t) | IRType.uint64 => pure ("lean_unbox_uint64", ← toLLVMType t) + | IRType.isize => pure ("lean_unbox_isize", ← toLLVMType t) + | IRType.int32 => pure ("lean_unbox_int32", ← toLLVMType t) + | IRType.int64 => pure ("lean_unbox_int64", ← toLLVMType t) | IRType.float => pure ("lean_unbox_float", ← toLLVMType t) | IRType.float32 => pure ("lean_unbox_float32", ← toLLVMType t) | _ => pure ("lean_unbox", ← LLVM.size_tType llvmctx) @@ -979,7 +998,7 @@ def emitVDecl (builder : LLVM.Builder llvmctx) (z : VarId) (t : IRType) (v : Exp | 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.uproj i s x => emitUProj builder z i s 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 @@ -1022,8 +1041,9 @@ def emitSet (builder : LLVM.Builder llvmctx) (x : VarId) (i : Nat) (y : Arg) : M let fnty ← LLVM.functionType retty argtys let _ ← LLVM.buildCall2 builder fnty fn #[← emitLhsVal builder x, ← constIntUnsigned i, (← emitArgVal builder y).2] -def emitUSet (builder : LLVM.Builder llvmctx) (x : VarId) (i : Nat) (y : VarId) : M llvmctx Unit := do - let fnName := "lean_ctor_set_usize" +def emitUSet (builder : LLVM.Builder llvmctx) (x : VarId) (i : Nat) (s : Bool) + (y : VarId) : M llvmctx Unit := do + let fnName := if s then "lean_ctor_set_isize" else "lean_ctor_set_usize" let retty ← LLVM.voidType llvmctx let argtys := #[ ← LLVM.voidPtrType llvmctx, ← LLVM.unsignedType llvmctx, ← LLVM.size_tType llvmctx] let fn ← getOrCreateFunctionPrototype (← getLLVMModule) retty fnName argtys @@ -1064,6 +1084,10 @@ def emitSSet (builder : LLVM.Builder llvmctx) (x : VarId) (n : Nat) (offset : Na | IRType.uint16 => pure ("lean_ctor_set_uint16", ← LLVM.i16Type llvmctx) | IRType.uint32 => pure ("lean_ctor_set_uint32", ← LLVM.i32Type llvmctx) | IRType.uint64 => pure ("lean_ctor_set_uint64", ← LLVM.i64Type llvmctx) + | IRType.int8 => pure ("lean_ctor_set_int8", ← LLVM.i8Type llvmctx) + | IRType.int16 => pure ("lean_ctor_set_int16", ← LLVM.i16Type llvmctx) + | IRType.int32 => pure ("lean_ctor_set_int32", ← LLVM.i32Type llvmctx) + | IRType.int64 => pure ("lean_ctor_set_int64", ← LLVM.i64Type llvmctx) | _ => throw s!"invalid type for 'lean_ctor_set': '{t}'" let argtys := #[ ← LLVM.voidPtrType llvmctx, ← LLVM.unsignedType llvmctx, setty] let retty ← LLVM.voidType llvmctx @@ -1168,7 +1192,7 @@ 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.uset x i y s b => emitUSet builder x i s 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" diff --git a/src/Lean/Compiler/IR/ExpandResetReuse.lean b/src/Lean/Compiler/IR/ExpandResetReuse.lean index dd8c08073810..b537f8ca9e9e 100644 --- a/src/Lean/Compiler/IR/ExpandResetReuse.lean +++ b/src/Lean/Compiler/IR/ExpandResetReuse.lean @@ -63,7 +63,7 @@ partial def eraseProjIncForAux (y : VarId) (bs : Array FnBody) (mask : Mask) (ke let b := bs.back! match b with | .vdecl _ _ (.sproj _ _ _) _ => keepInstr b - | .vdecl _ _ (.uproj _ _) _ => keepInstr b + | .vdecl _ _ (.uproj _ _ _) _ => keepInstr b | .inc z n c p _ => if n == 0 then done () else let b' := bs[bs.size - 2] @@ -160,7 +160,7 @@ def isSelfSet (ctx : Context) (x : VarId) (i : Nat) (y : Arg) : Bool := /-- 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 + | some (Expr.uproj j _ w) => j == i && w == x | _ => false /-- Given `sset x[n, i] := y`, return true iff `y := sproj[n, i] x` -/ @@ -174,9 +174,9 @@ 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 i y s b => if isSelfUSet ctx x i y then removeSelfSet ctx b - else FnBody.uset x i y (removeSelfSet ctx b) + else FnBody.uset x i y s (removeSelfSet ctx b) | FnBody.sset x 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) diff --git a/src/Lean/Compiler/IR/Format.lean b/src/Lean/Compiler/IR/Format.lean index ff1b04a80271..2409acf81531 100644 --- a/src/Lean/Compiler/IR/Format.lean +++ b/src/Lean/Compiler/IR/Format.lean @@ -44,7 +44,7 @@ private def formatExpr : Expr → Format | 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.uproj i s x => "uproj[" ++ format i ++ (if s then ", s" else ", u") ++ "] " ++ format x | Expr.sproj n o x => "sproj[" ++ format n ++ ", " ++ format o ++ "] " ++ format x | Expr.fap c ys => format c ++ formatArray ys | Expr.pap c ys => "pap " ++ format c ++ formatArray ys @@ -65,6 +65,11 @@ private partial def formatIRType : IRType → Format | IRType.uint32 => "u32" | IRType.uint64 => "u64" | IRType.usize => "usize" + | IRType.int8 => "i8" + | IRType.int16 => "i16" + | IRType.int32 => "i32" + | IRType.int64 => "i64" + | IRType.isize => "isize" | IRType.erased => "◾" | IRType.void => "void" | IRType.object => "obj" @@ -96,7 +101,7 @@ 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.uset x i y s _ => "uset " ++ format x ++ "[" ++ format i ++ (if s then ", s" else ", u") ++ "] := " ++ format y | FnBody.sset x i o y ty _ => "sset " ++ format x ++ "[" ++ 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 @@ -116,7 +121,7 @@ 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.uset x i y s b => "uset " ++ format x ++ "[" ++ format i ++ (if s then ", s" else ", u") ++ "] := " ++ 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.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 diff --git a/src/Lean/Compiler/IR/FreeVars.lean b/src/Lean/Compiler/IR/FreeVars.lean index 05fab205f123..82c83cc64f8b 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 10a25e591145..ed69f383b62b 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..a02dd4d32f67 100644 --- a/src/Lean/Compiler/IR/NormIds.lean +++ b/src/Lean/Compiler/IR/NormIds.lean @@ -65,7 +65,7 @@ def normExpr : Expr → M Expr | 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.uproj i s x, m => Expr.uproj i s (normVar x m) | Expr.sproj n o x, m => Expr.sproj 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) @@ -101,7 +101,7 @@ partial def normFnBody : FnBody → N FnBody 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.uset x i y s b => return FnBody.uset (← normVar x) i (← normVar y) s (← 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) @@ -142,7 +142,7 @@ def mapExpr (f : VarId → VarId) : Expr → Expr | 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.uproj i s x => Expr.uproj i s (f x) | Expr.sproj n o x => Expr.sproj n o (f x) | Expr.fap c ys => Expr.fap c (mapArgs f ys) | Expr.pap c ys => Expr.pap c (mapArgs f ys) @@ -157,7 +157,7 @@ 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.uset x i y s b => FnBody.uset (f x) i (f y) s (mapFnBody f b) | FnBody.sset x i o y t b => FnBody.sset (f x) 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) 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..265f19eed7dc 100644 --- a/src/Lean/Compiler/IR/RC.lean +++ b/src/Lean/Compiler/IR/RC.lean @@ -185,7 +185,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 @@ -360,7 +360,7 @@ private def processVDecl (ctx : Context) (z : VarId) (t : IRType) (v : Expr) (b let b := addDecIfNeeded ctx x b bLiveVars let b := if !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 | .unbox x => .vdecl z t v (addDecIfNeeded ctx x b bLiveVars) | .fap f ys => let ps := (getDecl ctx f).params @@ -406,11 +406,11 @@ 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 i y signed 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⟩ + ⟨.uset x i y signed b, s⟩ | .sset x 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 diff --git a/src/Lean/Compiler/IR/ToIR.lean b/src/Lean/Compiler/IR/ToIR.lean index 75ca86f7a4f4..95f63a01881b 100644 --- a/src/Lean/Compiler/IR/ToIR.lean +++ b/src/Lean/Compiler/IR/ToIR.lean @@ -108,7 +108,7 @@ 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⟩ + | .usize i s => ⟨.expr (.uproj i s base), IRType.ptrSizedTypeForSign s⟩ | .scalar _ offset irType => ⟨.expr (.sproj (ctorInfo.size + ctorInfo.usize) offset base), irType⟩ | .erased => ⟨.erased, .erased⟩ | .void => ⟨.erased, .void⟩ @@ -228,9 +228,9 @@ partial def lowerLet (decl : LCNF.LetDecl) (k : LCNF.Code) : M FnBody := do match irArgs[i]? with | some (.var varId) => match fields[i]! with - | .usize usizeIdx => + | .usize usizeIdx s => let k ← loop (i + 1) - return .uset objVar usizeIdx varId k + return .uset objVar usizeIdx varId s k | .scalar _ offset argType => let k ← loop (i + 1) return .sset objVar (ctorInfo.size + ctorInfo.usize) offset varId argType k diff --git a/src/Lean/Compiler/IR/ToIRType.lean b/src/Lean/Compiler/IR/ToIRType.lean index 602c4509e17f..b3d697661cbe 100644 --- a/src/Lean/Compiler/IR/ToIRType.lean +++ b/src/Lean/Compiler/IR/ToIRType.lean @@ -64,6 +64,11 @@ where fillCache : CoreM IRType := do | ``UInt32 => return .uint32 | ``UInt64 => return .uint64 | ``USize => return .usize + | ``Int8 => return .int8 + | ``Int16 => return .int16 + | ``Int32 => return .int32 + | ``Int64 => return .int64 + | ``ISize => return .isize | ``Float => return .float | ``Float32 => return .float32 | ``lcErased => return .erased @@ -132,7 +137,7 @@ where inductive CtorFieldInfo where | erased | object (i : Nat) (type : IRType) - | usize (i : Nat) + | usize (i : Nat) (signed : Bool) | scalar (sz : Nat) (offset : Nat) (type : IRType) | void deriving Inhabited @@ -143,7 +148,7 @@ def format : CtorFieldInfo → Format | erased => "◾" | void => "void" | object i type => f!"obj@{i}:{type}" - | usize i => f!"usize@{i}" + | usize i s => if s then f!"isize@{i}" else f!"usize@{i}" | scalar sz offset type => f!"scalar#{sz}@{offset}:{type}" instance : ToFormat CtorFieldInfo := ⟨format⟩ @@ -184,21 +189,34 @@ where fillCache := do let i := nextIdx nextIdx := nextIdx + 1 pure <| .object i irFieldType - | .usize => pure <| .usize 0 + | .usize => pure <| .usize 0 false + | .isize => pure <| .usize 0 true | .erased => .pure <| .erased | .void => .pure <| .void | .uint8 => has1BScalar := true .pure <| .scalar 1 0 .uint8 + | .int8 => + has1BScalar := true + .pure <| .scalar 1 0 .int8 | .uint16 => has2BScalar := true .pure <| .scalar 2 0 .uint16 + | .int16 => + has2BScalar := true + .pure <| .scalar 2 0 .int16 | .uint32 => has4BScalar := true .pure <| .scalar 4 0 .uint32 + | .int32 => + has4BScalar := true + .pure <| .scalar 4 0 .int32 | .uint64 => has8BScalar := true .pure <| .scalar 8 0 .uint64 + | .int64 => + has8BScalar := true + .pure <| .scalar 8 0 .int64 | .float32 => has4BScalar := true .pure <| .scalar 4 0 .float32 @@ -210,9 +228,9 @@ where fillCache := do let numObjs := nextIdx ⟨fields, nextIdx⟩ := Id.run <| StateT.run (s := nextIdx) <| fields.mapM fun field => do match field with - | .usize _ => do + | .usize _ s => do let i ← modifyGet fun nextIdx => (nextIdx, nextIdx + 1) - return .usize i + return .usize i s | .object .. | .scalar .. | .erased | .void => return field let numUSize := nextIdx - numObjs let adjustScalarsForSize (fields : Array CtorFieldInfo) (size : Nat) (nextOffset : Nat) @@ -225,7 +243,7 @@ where fillCache := do return .scalar sz offset type else return field - | .object .. | .usize _ | .erased | .void => return field + | .object .. | .usize _ _ | .erased | .void => return field let mut nextOffset := 0 if has8BScalar then ⟨fields, nextOffset⟩ := adjustScalarsForSize fields 8 nextOffset diff --git a/src/Lean/Compiler/LCNF/ToMono.lean b/src/Lean/Compiler/LCNF/ToMono.lean index 5d460dabd568..1a09d4b25714 100644 --- a/src/Lean/Compiler/LCNF/ToMono.lean +++ b/src/Lean/Compiler/LCNF/ToMono.lean @@ -264,6 +264,17 @@ partial def casesUIntToMono (c : Cases) (uintName : Name) (_ : c.typeName == uin let k ← k.toMono return .let decl k +/-- Eliminate `cases` for `SInt` types. -/ +partial def casesSIntToMono (c : Cases) (sintName : Name) (sintProj : String) (_ : c.typeName == sintName) : ToMonoM Code := do + assert! c.alts.size == 1 + let .alt _ ps k := c.alts[0]! | unreachable! + eraseParams ps + let p := ps[0]! + let decl := { fvarId := p.fvarId, binderName := p.binderName, type := anyExpr, value := .const (.str sintName sintProj) [] #[.fvar c.discr] } + modifyLCtx fun lctx => lctx.addLetDecl decl + let k ← k.toMono + return .let decl k + /-- Eliminate `cases` for `Array. -/ partial def casesArrayToMono (c : Cases) (_ : c.typeName == ``Array) : ToMonoM Code := do assert! c.alts.size == 1 @@ -386,6 +397,14 @@ partial def Code.toMono (code : Code) : ToMonoM Code := do casesUIntToMono c ``UInt32 h else if h : c.typeName == ``UInt64 then casesUIntToMono c ``UInt64 h + else if h : c.typeName == ``Int8 then + casesSIntToMono c ``Int8 "toUInt8" h + else if h : c.typeName == ``Int16 then + casesSIntToMono c ``Int16 "toUInt16" h + else if h : c.typeName == ``Int32 then + casesSIntToMono c ``Int32 "toUInt32" h + else if h : c.typeName == ``Int64 then + casesSIntToMono c ``Int64 "toUInt64" h else if h : c.typeName == ``Array then casesArrayToMono c h else if h : c.typeName == ``ByteArray then diff --git a/src/Lean/Compiler/LCNF/Util.lean b/src/Lean/Compiler/LCNF/Util.lean index 041dceb5df1d..9b68f01f85f5 100644 --- a/src/Lean/Compiler/LCNF/Util.lean +++ b/src/Lean/Compiler/LCNF/Util.lean @@ -83,6 +83,7 @@ List of types that have builtin runtime support def builtinRuntimeTypes : Array Name := #[ ``String, ``UInt8, ``UInt16, ``UInt32, ``UInt64, ``USize, + ``Int8, ``Int16, ``Int32, ``Int64, ``ISize, ``Float, ``Float32, ``Thunk, ``Task, ``Array, ``ByteArray, ``FloatArray, diff --git a/src/include/lean/lean.h b/src/include/lean/lean.h index b3ae906c7935..8c13940867aa 100644 --- a/src/include/lean/lean.h +++ b/src/include/lean/lean.h @@ -641,26 +641,51 @@ static inline size_t lean_ctor_get_usize(b_lean_obj_arg o, unsigned i) { return *((size_t*)(lean_ctor_obj_cptr(o) + i)); } +static inline ptrdiff_t lean_ctor_get_isize(b_lean_obj_arg o, unsigned i) { + assert(i >= lean_ctor_num_objs(o)); + return *((ptrdiff_t*)(lean_ctor_obj_cptr(o) + i)); +} + static inline uint8_t lean_ctor_get_uint8(b_lean_obj_arg o, unsigned offset) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); return *((uint8_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); } +static inline int8_t lean_ctor_get_int8(b_lean_obj_arg o, unsigned offset) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + return *((int8_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); +} + static inline uint16_t lean_ctor_get_uint16(b_lean_obj_arg o, unsigned offset) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); return *((uint16_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); } +static inline int16_t lean_ctor_get_int16(b_lean_obj_arg o, unsigned offset) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + return *((int16_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); +} + static inline uint32_t lean_ctor_get_uint32(b_lean_obj_arg o, unsigned offset) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); return *((uint32_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); } +static inline int32_t lean_ctor_get_int32(b_lean_obj_arg o, unsigned offset) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + return *((int32_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); +} + static inline uint64_t lean_ctor_get_uint64(b_lean_obj_arg o, unsigned offset) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); return *((uint64_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); } +static inline int64_t lean_ctor_get_int64(b_lean_obj_arg o, unsigned offset) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + return *((int64_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); +} + static inline double lean_ctor_get_float(b_lean_obj_arg o, unsigned offset) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); return *((double*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)); @@ -676,26 +701,51 @@ static inline void lean_ctor_set_usize(b_lean_obj_arg o, unsigned i, size_t v) { *((size_t*)(lean_ctor_obj_cptr(o) + i)) = v; } +static inline void lean_ctor_set_isize(b_lean_obj_arg o, unsigned i, ptrdiff_t v) { + assert(i >= lean_ctor_num_objs(o)); + *((ptrdiff_t*)(lean_ctor_obj_cptr(o) + i)) = v; +} + static inline void lean_ctor_set_uint8(b_lean_obj_arg o, unsigned offset, uint8_t v) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); *((uint8_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; } +static inline void lean_ctor_set_int8(b_lean_obj_arg o, unsigned offset, int8_t v) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + *((int8_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; +} + static inline void lean_ctor_set_uint16(b_lean_obj_arg o, unsigned offset, uint16_t v) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); *((uint16_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; } +static inline void lean_ctor_set_int16(b_lean_obj_arg o, unsigned offset, int16_t v) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + *((int16_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; +} + static inline void lean_ctor_set_uint32(b_lean_obj_arg o, unsigned offset, uint32_t v) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); *((uint32_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; } +static inline void lean_ctor_set_int32(b_lean_obj_arg o, unsigned offset, int32_t v) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + *((int32_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; +} + static inline void lean_ctor_set_uint64(b_lean_obj_arg o, unsigned offset, uint64_t v) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); *((uint64_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; } +static inline void lean_ctor_set_int64(b_lean_obj_arg o, unsigned offset, int64_t v) { + assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); + *((int64_t*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; +} + static inline void lean_ctor_set_float(b_lean_obj_arg o, unsigned offset, double v) { assert(offset >= lean_ctor_num_objs(o) * sizeof(void*)); *((double*)((uint8_t*)(lean_ctor_obj_cptr(o)) + offset)) = v; @@ -1783,11 +1833,11 @@ static inline uint16_t lean_bool_to_uint16(uint8_t a) { return (uint16_t)a; } static inline uint32_t lean_bool_to_uint32(uint8_t a) { return (uint32_t)a; } static inline uint64_t lean_bool_to_uint64(uint8_t a) { return (uint64_t)a; } static inline size_t lean_bool_to_usize(uint8_t a) { return (size_t)a; } -static inline uint8_t lean_bool_to_int8(uint8_t a) { return (uint8_t)(int8_t)a; } -static inline uint16_t lean_bool_to_int16(uint8_t a) { return (uint16_t)(int16_t)a; } -static inline uint32_t lean_bool_to_int32(uint8_t a) { return (uint32_t)(int32_t)a; } -static inline uint64_t lean_bool_to_int64(uint8_t a) { return (uint64_t)(int64_t)a; } -static inline size_t lean_bool_to_isize(uint8_t a) { return (size_t)(ptrdiff_t)a; } +static inline int8_t lean_bool_to_int8(uint8_t a) { return (int8_t)a; } +static inline int16_t lean_bool_to_int16(uint8_t a) { return (int16_t)a; } +static inline int32_t lean_bool_to_int32(uint8_t a) { return (int32_t)a; } +static inline int64_t lean_bool_to_int64(uint8_t a) { return (int64_t)a; } +static inline ptrdiff_t lean_bool_to_isize(uint8_t a) { return (ptrdiff_t)a; } /* UInt8 */ @@ -1993,7 +2043,7 @@ static inline uint64_t lean_usize_to_uint64(size_t a) { return ((uint64_t)a); } /* Int8 */ LEAN_EXPORT int8_t lean_int8_of_big_int(b_lean_obj_arg a); -static inline uint8_t lean_int8_of_int(b_lean_obj_arg a) { +static inline int8_t lean_int8_of_int(b_lean_obj_arg a) { int8_t res; if (lean_is_scalar(a)) { @@ -2002,10 +2052,10 @@ static inline uint8_t lean_int8_of_int(b_lean_obj_arg a) { res = lean_int8_of_big_int(a); } - return (uint8_t)res; + return res; } -static inline uint8_t lean_int8_of_nat(b_lean_obj_arg a) { +static inline int8_t lean_int8_of_nat(b_lean_obj_arg a) { int8_t res; if (lean_is_scalar(a)) { @@ -2014,135 +2064,103 @@ static inline uint8_t lean_int8_of_nat(b_lean_obj_arg a) { res = lean_int8_of_big_int(a); } - return (uint8_t)res; + return res; } -static inline lean_obj_res lean_int8_to_int(uint8_t a) { - int8_t arg = (int8_t)a; - return lean_int64_to_int((int64_t)arg); +static inline int8_t lean_int8_of_uint8_mk(uint8_t a) { + return (int8_t)a; } -static inline uint8_t lean_int8_neg(uint8_t a) { - int8_t arg = (int8_t)a; - - return (uint8_t)(-arg); +static inline uint8_t lean_int8_to_uint8(int8_t a) { + return (uint8_t)a; } -static inline uint8_t lean_int8_add(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs + rhs); +static inline lean_obj_res lean_int8_to_int(int8_t a) { + return lean_int64_to_int((int64_t)a); } -static inline uint8_t lean_int8_sub(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs - rhs); +static inline int8_t lean_int8_neg(int8_t a) { + return -a; } -static inline uint8_t lean_int8_mul(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs * rhs); +static inline int8_t lean_int8_add(int8_t lhs, int8_t rhs) { + return lhs + rhs; } -static inline uint8_t lean_int8_div(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(rhs == 0 ? 0 : lhs / rhs); +static inline int8_t lean_int8_sub(int8_t lhs, int8_t rhs) { + return lhs - rhs; } -static inline uint8_t lean_int8_mod(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(rhs == 0 ? lhs : lhs % rhs); +static inline int8_t lean_int8_mul(int8_t lhs, int8_t rhs) { + return lhs * rhs; } -static inline uint8_t lean_int8_land(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs & rhs); +static inline int8_t lean_int8_div(int8_t lhs, int8_t rhs) { + return rhs == 0 ? 0 : lhs / rhs; } -static inline uint8_t lean_int8_lor(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs | rhs); +static inline int8_t lean_int8_mod(int8_t lhs, int8_t rhs) { + return rhs == 0 ? lhs : lhs % rhs; } -static inline uint8_t lean_int8_xor(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - - return (uint8_t)(lhs ^ rhs); +static inline int8_t lean_int8_land(int8_t lhs, int8_t rhs) { + return lhs & rhs; } -static inline uint8_t lean_int8_shift_right(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (((int8_t)a2 % 8) + 8) % 8; // this is smod 8 +static inline int8_t lean_int8_lor(int8_t lhs, int8_t rhs) { + return lhs | rhs; +} - return (uint8_t)(lhs >> rhs); +static inline int8_t lean_int8_xor(int8_t lhs, int8_t rhs) { + return lhs ^ rhs; } -static inline uint8_t lean_int8_shift_left(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (((int8_t)a2 % 8) + 8) % 8; // this is smod 8 +static inline int8_t lean_int8_shift_right(int8_t a1, int8_t a2) { + int8_t lhs = a1; + int8_t rhs = ((a2 % 8) + 8) % 8; // this is smod 8 - return (uint8_t)(lhs << rhs); + return lhs >> rhs; } -static inline uint8_t lean_int8_complement(uint8_t a) { - int8_t arg = (int8_t)a; +static inline int8_t lean_int8_shift_left(int8_t a1, int8_t a2) { + int8_t lhs = a1; + int8_t rhs = ((a2 % 8) + 8) % 8; // this is smod 8 - return (uint8_t)(~arg); + return lhs << rhs; } -static inline uint8_t lean_int8_abs(uint8_t a) { - int8_t arg = (int8_t)a; +static inline int8_t lean_int8_complement(int8_t a) { + return ~a; +} +static inline int8_t lean_int8_abs(int8_t a) { // Recall that we are compiling with -fwrapv so this is guaranteed to // map INT8_MIN to INT8_MIN - return (uint8_t)(arg < 0 ? -arg : arg); + return (a < 0 ? -a : a); } -static inline uint8_t lean_int8_dec_eq(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - +static inline uint8_t lean_int8_dec_eq(int8_t lhs, int8_t rhs) { return lhs == rhs; } -static inline uint8_t lean_int8_dec_lt(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - +static inline uint8_t lean_int8_dec_lt(int8_t lhs, int8_t rhs) { return lhs < rhs; } -static inline uint8_t lean_int8_dec_le(uint8_t a1, uint8_t a2) { - int8_t lhs = (int8_t)a1; - int8_t rhs = (int8_t)a2; - +static inline uint8_t lean_int8_dec_le(int8_t lhs, int8_t rhs) { return lhs <= rhs; } /* Int8 -> other */ -static inline uint16_t lean_int8_to_int16(uint8_t a) { return (uint16_t)(int16_t)(int8_t)a; } -static inline uint32_t lean_int8_to_int32(uint8_t a) { return (uint32_t)(int32_t)(int8_t)a; } -static inline uint64_t lean_int8_to_int64(uint8_t a) { return (uint64_t)(int64_t)(int8_t)a; } -static inline size_t lean_int8_to_isize(uint8_t a) { return (size_t)(ptrdiff_t)(int8_t)a; } +static inline int16_t lean_int8_to_int16(int8_t a) { return (int16_t)a; } +static inline int32_t lean_int8_to_int32(int8_t a) { return (int32_t)a; } +static inline int64_t lean_int8_to_int64(int8_t a) { return (int64_t)a; } +static inline ptrdiff_t lean_int8_to_isize(int8_t a) { return (ptrdiff_t)a; } /* Int16 */ LEAN_EXPORT int16_t lean_int16_of_big_int(b_lean_obj_arg a); -static inline uint16_t lean_int16_of_int(b_lean_obj_arg a) { +static inline int16_t lean_int16_of_int(b_lean_obj_arg a) { int16_t res; if (lean_is_scalar(a)) { @@ -2151,10 +2169,10 @@ static inline uint16_t lean_int16_of_int(b_lean_obj_arg a) { res = lean_int16_of_big_int(a); } - return (uint16_t)res; + return res; } -static inline uint16_t lean_int16_of_nat(b_lean_obj_arg a) { +static inline int16_t lean_int16_of_nat(b_lean_obj_arg a) { int16_t res; if (lean_is_scalar(a)) { @@ -2163,134 +2181,103 @@ static inline uint16_t lean_int16_of_nat(b_lean_obj_arg a) { res = lean_int16_of_big_int(a); } - return (uint16_t)res; + return res; } -static inline lean_obj_res lean_int16_to_int(uint16_t a) { - int16_t arg = (int16_t)a; - return lean_int64_to_int((int64_t)arg); +static inline int16_t lean_int16_of_uint16_mk(uint16_t a) { + return (int16_t)a; } -static inline uint16_t lean_int16_neg(uint16_t a) { - int16_t arg = (int16_t)a; - - return (uint16_t)(-arg); +static inline uint16_t lean_int16_to_uint16(int16_t a) { + return (uint16_t)a; } -static inline uint16_t lean_int16_add(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs + rhs); +static inline lean_obj_res lean_int16_to_int(int16_t a) { + int16_t arg = (int16_t)a; + return lean_int64_to_int((int64_t)arg); } -static inline uint16_t lean_int16_sub(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs - rhs); +static inline int16_t lean_int16_neg(int16_t a) { + return -a; } -static inline uint16_t lean_int16_mul(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs * rhs); +static inline int16_t lean_int16_add(int16_t lhs, int16_t rhs) { + return lhs + rhs; } -static inline uint16_t lean_int16_div(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(rhs == 0 ? 0 : lhs / rhs); +static inline int16_t lean_int16_sub(int16_t lhs, int16_t rhs) { + return lhs - rhs; } -static inline uint16_t lean_int16_mod(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(rhs == 0 ? lhs : lhs % rhs); +static inline int16_t lean_int16_mul(int16_t lhs, int16_t rhs) { + return lhs * rhs; } -static inline uint16_t lean_int16_land(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs & rhs); +static inline int16_t lean_int16_div(int16_t lhs, int16_t rhs) { + return rhs == 0 ? 0 : lhs / rhs; } -static inline uint16_t lean_int16_lor(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs | rhs); +static inline int16_t lean_int16_mod(int16_t lhs, int16_t rhs) { + return rhs == 0 ? lhs : lhs % rhs; } -static inline uint16_t lean_int16_xor(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - - return (uint16_t)(lhs ^ rhs); +static inline int16_t lean_int16_land(int16_t lhs, int16_t rhs) { + return lhs & rhs; } -static inline uint16_t lean_int16_shift_right(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (((int16_t)a2 % 16) + 16) % 16; // this is smod 16 +static inline int16_t lean_int16_lor(int16_t lhs, int16_t rhs) { + return lhs | rhs; +} - return (uint16_t)(lhs >> rhs); +static inline int16_t lean_int16_xor(int16_t lhs, int16_t rhs) { + return lhs ^ rhs; } -static inline uint16_t lean_int16_shift_left(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (((int16_t)a2 % 16) + 16) % 16; // this is smod 16 +static inline int16_t lean_int16_shift_right(int16_t a1, int16_t a2) { + int16_t lhs = a1; + int16_t rhs = ((a2 % 16) + 16) % 16; // this is smod 16 - return (uint16_t)(lhs << rhs); + return lhs >> rhs; } -static inline uint16_t lean_int16_complement(uint16_t a) { - int16_t arg = (int16_t)a; +static inline int16_t lean_int16_shift_left(int16_t a1, int16_t a2) { + int16_t lhs = a1; + int16_t rhs = ((a2 % 16) + 16) % 16; // this is smod 16 - return (uint16_t)(~arg); + return lhs << rhs; } -static inline uint16_t lean_int16_abs(uint16_t a) { - int16_t arg = (int16_t)a; +static inline int16_t lean_int16_complement(int16_t a) { + return ~a; +} +static inline int16_t lean_int16_abs(int16_t a) { // Recall that we are compiling with -fwrapv so this is guaranteed to // map INT16_MIN to INT16_MIN - return (uint16_t)(arg < 0 ? -arg : arg); + return (a < 0 ? -a : a); } -static inline uint8_t lean_int16_dec_eq(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - +static inline uint8_t lean_int16_dec_eq(int16_t lhs, int16_t rhs) { return lhs == rhs; } -static inline uint8_t lean_int16_dec_lt(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - +static inline uint8_t lean_int16_dec_lt(int16_t lhs, int16_t rhs) { return lhs < rhs; } -static inline uint8_t lean_int16_dec_le(uint16_t a1, uint16_t a2) { - int16_t lhs = (int16_t)a1; - int16_t rhs = (int16_t)a2; - +static inline uint8_t lean_int16_dec_le(int16_t lhs, int16_t rhs) { return lhs <= rhs; } /* Int16 -> other */ -static inline uint8_t lean_int16_to_int8(uint16_t a) { return (uint8_t)(int8_t)(int16_t)a; } -static inline uint32_t lean_int16_to_int32(uint16_t a) { return (uint32_t)(int32_t)(int16_t)a; } -static inline uint64_t lean_int16_to_int64(uint16_t a) { return (uint64_t)(int64_t)(int16_t)a; } -static inline size_t lean_int16_to_isize(uint16_t a) { return (size_t)(ptrdiff_t)(int16_t)a; } +static inline int8_t lean_int16_to_int8(int16_t a) { return (int8_t)a; } +static inline int32_t lean_int16_to_int32(int16_t a) { return (int32_t)a; } +static inline int64_t lean_int16_to_int64(int16_t a) { return (int64_t)a; } +static inline ptrdiff_t lean_int16_to_isize(int16_t a) { return (ptrdiff_t)a; } /* Int32 */ LEAN_EXPORT int32_t lean_int32_of_big_int(b_lean_obj_arg a); -static inline uint32_t lean_int32_of_int(b_lean_obj_arg a) { +static inline int32_t lean_int32_of_int(b_lean_obj_arg a) { int32_t res; if (lean_is_scalar(a)) { @@ -2299,10 +2286,10 @@ static inline uint32_t lean_int32_of_int(b_lean_obj_arg a) { res = lean_int32_of_big_int(a); } - return (uint32_t)res; + return res; } -static inline uint32_t lean_int32_of_nat(b_lean_obj_arg a) { +static inline int32_t lean_int32_of_nat(b_lean_obj_arg a) { int32_t res; if (lean_is_scalar(a)) { @@ -2311,134 +2298,102 @@ static inline uint32_t lean_int32_of_nat(b_lean_obj_arg a) { res = lean_int32_of_big_int(a); } - return (uint32_t)res; + return res; } -static inline lean_obj_res lean_int32_to_int(uint32_t a) { - int32_t arg = (int32_t)a; - return lean_int64_to_int((int64_t)arg); +static inline int32_t lean_int32_of_uint32_mk(uint32_t a) { + return (int32_t)a; } -static inline uint32_t lean_int32_neg(uint32_t a) { - int32_t arg = (int32_t)a; - - return (uint32_t)(-arg); +static inline uint32_t lean_int32_to_uint32(int32_t a) { + return (uint32_t)a; } -static inline uint32_t lean_int32_add(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs + rhs); +static inline lean_obj_res lean_int32_to_int(int32_t a) { + return lean_int64_to_int((int64_t)a); } -static inline uint32_t lean_int32_sub(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs - rhs); +static inline int32_t lean_int32_neg(int32_t a) { + return -a; } -static inline uint32_t lean_int32_mul(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs * rhs); +static inline int32_t lean_int32_add(int32_t lhs, int32_t rhs) { + return lhs + rhs; } -static inline uint32_t lean_int32_div(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(rhs == 0 ? 0 : lhs / rhs); +static inline int32_t lean_int32_sub(int32_t lhs, int32_t rhs) { + return lhs - rhs; } -static inline uint32_t lean_int32_mod(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(rhs == 0 ? lhs : lhs % rhs); +static inline int32_t lean_int32_mul(int32_t lhs, int32_t rhs) { + return lhs * rhs; } -static inline uint32_t lean_int32_land(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs & rhs); +static inline int32_t lean_int32_div(int32_t lhs, int32_t rhs) { + return rhs == 0 ? 0 : lhs / rhs; } -static inline uint32_t lean_int32_lor(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs | rhs); +static inline int32_t lean_int32_mod(int32_t lhs, int32_t rhs) { + return rhs == 0 ? lhs : lhs % rhs; } -static inline uint32_t lean_int32_xor(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - - return (uint32_t)(lhs ^ rhs); +static inline int32_t lean_int32_land(int32_t lhs, int32_t rhs) { + return lhs & rhs; } -static inline uint32_t lean_int32_shift_right(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (((int32_t)a2 % 32) + 32) % 32; // this is smod 32 +static inline int32_t lean_int32_lor(int32_t lhs, int32_t rhs) { + return lhs | rhs; +} - return (uint32_t)(lhs >> rhs); +static inline int32_t lean_int32_xor(int32_t lhs, int32_t rhs) { + return lhs ^ rhs; } -static inline uint32_t lean_int32_shift_left(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (((int32_t)a2 % 32) + 32) % 32; // this is smod 32 +static inline int32_t lean_int32_shift_right(int32_t a1, int32_t a2) { + int32_t lhs = a1; + int32_t rhs = ((a2 % 32) + 32) % 32; // this is smod 32 - return (uint32_t)(lhs << rhs); + return lhs >> rhs; } -static inline uint32_t lean_int32_complement(uint32_t a) { - int32_t arg = (int32_t)a; +static inline int32_t lean_int32_shift_left(int32_t a1, int32_t a2) { + int32_t lhs = a1; + int32_t rhs = ((a2 % 32) + 32) % 32; // this is smod 32 - return (uint32_t)(~arg); + return lhs << rhs; } -static inline uint32_t lean_int32_abs(uint32_t a) { - int32_t arg = (int32_t)a; +static inline int32_t lean_int32_complement(int32_t a) { + return ~a; +} +static inline int32_t lean_int32_abs(int32_t a) { // Recall that we are compiling with -fwrapv so this is guaranteed to // map INT32_MIN to INT32_MIN - return (uint32_t)(arg < 0 ? -arg : arg); + return (a < 0 ? -a : a); } -static inline uint8_t lean_int32_dec_eq(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - +static inline uint8_t lean_int32_dec_eq(int32_t lhs, int32_t rhs) { return lhs == rhs; } -static inline uint8_t lean_int32_dec_lt(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - +static inline uint8_t lean_int32_dec_lt(int32_t lhs, int32_t rhs) { return lhs < rhs; } -static inline uint8_t lean_int32_dec_le(uint32_t a1, uint32_t a2) { - int32_t lhs = (int32_t)a1; - int32_t rhs = (int32_t)a2; - +static inline uint8_t lean_int32_dec_le(int32_t lhs, int32_t rhs) { return lhs <= rhs; } /* Int32 -> other */ -static inline uint8_t lean_int32_to_int8(uint32_t a) { return (uint8_t)(int8_t)(int32_t)a; } -static inline uint16_t lean_int32_to_int16(uint32_t a) { return (uint16_t)(int16_t)(int32_t)a; } -static inline uint64_t lean_int32_to_int64(uint32_t a) { return (uint64_t)(int64_t)(int32_t)a; } -static inline size_t lean_int32_to_isize(uint32_t a) { return (size_t)(ptrdiff_t)(int32_t)a; } +static inline int8_t lean_int32_to_int8(int32_t a) { return (int8_t)a; } +static inline int16_t lean_int32_to_int16(int32_t a) { return (int16_t)a; } +static inline int64_t lean_int32_to_int64(int32_t a) { return (int64_t)a; } +static inline ptrdiff_t lean_int32_to_isize(int32_t a) { return (ptrdiff_t)a; } /* Int64 */ LEAN_EXPORT int64_t lean_int64_of_big_int(b_lean_obj_arg a); -static inline uint64_t lean_int64_of_int(b_lean_obj_arg a) { +static inline int64_t lean_int64_of_int(b_lean_obj_arg a) { int64_t res; if (lean_is_scalar(a)) { @@ -2447,10 +2402,10 @@ static inline uint64_t lean_int64_of_int(b_lean_obj_arg a) { res = lean_int64_of_big_int(a); } - return (uint64_t)res; + return res; } -static inline uint64_t lean_int64_of_nat(b_lean_obj_arg a) { +static inline int64_t lean_int64_of_nat(b_lean_obj_arg a) { int64_t res; if (lean_is_scalar(a)) { @@ -2459,134 +2414,102 @@ static inline uint64_t lean_int64_of_nat(b_lean_obj_arg a) { res = lean_int64_of_big_int(a); } - return (uint64_t)res; + return res; } -static inline lean_obj_res lean_int64_to_int_sint(uint64_t a) { - int64_t arg = (int64_t)a; - return lean_int64_to_int(arg); +static inline int64_t lean_int64_of_uint64_mk(uint64_t a) { + return (int64_t)a; } -static inline uint64_t lean_int64_neg(uint64_t a) { - int64_t arg = (int64_t)a; - - return (uint64_t)(-arg); +static inline uint64_t lean_int64_to_uint64(int64_t a) { + return (uint64_t)a; } -static inline uint64_t lean_int64_add(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs + rhs); +static inline lean_obj_res lean_int64_to_int_sint(int64_t a) { + return lean_int64_to_int(a); } -static inline uint64_t lean_int64_sub(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs - rhs); +static inline int64_t lean_int64_neg(int64_t a) { + return -a; } -static inline uint64_t lean_int64_mul(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs * rhs); +static inline int64_t lean_int64_add(int64_t lhs, int64_t rhs) { + return lhs + rhs; } -static inline uint64_t lean_int64_div(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(rhs == 0 ? 0 : lhs / rhs); +static inline int64_t lean_int64_sub(int64_t lhs, int64_t rhs) { + return lhs - rhs; } -static inline uint64_t lean_int64_mod(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(rhs == 0 ? lhs : lhs % rhs); +static inline int64_t lean_int64_mul(int64_t lhs, int64_t rhs) { + return lhs * rhs; } -static inline uint64_t lean_int64_land(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs & rhs); +static inline int64_t lean_int64_div(int64_t lhs, int64_t rhs) { + return rhs == 0 ? 0 : lhs / rhs; } -static inline uint64_t lean_int64_lor(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs | rhs); +static inline int64_t lean_int64_mod(int64_t lhs, int64_t rhs) { + return rhs == 0 ? lhs : lhs % rhs; } -static inline uint64_t lean_int64_xor(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - - return (uint64_t)(lhs ^ rhs); +static inline int64_t lean_int64_land(int64_t lhs, int64_t rhs) { + return lhs & rhs; } -static inline uint64_t lean_int64_shift_right(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (((int64_t)a2 % 64) + 64) % 64; // this is smod 64 +static inline int64_t lean_int64_lor(int64_t lhs, int64_t rhs) { + return lhs | rhs; +} - return (uint64_t)(lhs >> rhs); +static inline int64_t lean_int64_xor(int64_t lhs, int64_t rhs) { + return lhs ^ rhs; } -static inline uint64_t lean_int64_shift_left(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (((int64_t)a2 % 64) + 64) % 64; // this is smod 64 +static inline int64_t lean_int64_shift_right(int64_t a1, int64_t a2) { + int64_t lhs = a1; + int64_t rhs = ((a2 % 64) + 64) % 64; // this is smod 64 - return (uint64_t)(lhs << rhs); + return lhs >> rhs; } -static inline uint64_t lean_int64_complement(uint64_t a) { - int64_t arg = (int64_t)a; +static inline int64_t lean_int64_shift_left(int64_t a1, int64_t a2) { + int64_t lhs = a1; + int64_t rhs = ((a2 % 64) + 64) % 64; // this is smod 64 - return (uint64_t)(~arg); + return lhs << rhs; } -static inline uint64_t lean_int64_abs(uint64_t a) { - int64_t arg = (int64_t)a; +static inline int64_t lean_int64_complement(int64_t a) { + return ~a; +} +static inline int64_t lean_int64_abs(int64_t a) { // Recall that we are compiling with -fwrapv so this is guaranteed to // map INT64_MIN to INT64_MIN - return (uint64_t)(arg < 0 ? -arg : arg); + return a < 0 ? -a : a; } -static inline uint8_t lean_int64_dec_eq(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - +static inline uint8_t lean_int64_dec_eq(int64_t lhs, int64_t rhs) { return lhs == rhs; } -static inline uint8_t lean_int64_dec_lt(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - +static inline uint8_t lean_int64_dec_lt(int64_t lhs, int64_t rhs) { return lhs < rhs; } -static inline uint8_t lean_int64_dec_le(uint64_t a1, uint64_t a2) { - int64_t lhs = (int64_t)a1; - int64_t rhs = (int64_t)a2; - +static inline uint8_t lean_int64_dec_le(int64_t lhs, int64_t rhs) { return lhs <= rhs; } /* Int64 -> other */ -static inline uint8_t lean_int64_to_int8(uint64_t a) { return (uint8_t)(int8_t)(int64_t)a; } -static inline uint16_t lean_int64_to_int16(uint64_t a) { return (uint16_t)(int16_t)(int64_t)a; } -static inline uint32_t lean_int64_to_int32(uint64_t a) { return (uint32_t)(int32_t)(int64_t)a; } -static inline size_t lean_int64_to_isize(uint64_t a) { return (size_t)(ptrdiff_t)(int64_t)a; } +static inline int8_t lean_int64_to_int8(int64_t a) { return (int8_t)a; } +static inline int16_t lean_int64_to_int16(int64_t a) { return (int16_t)a; } +static inline int32_t lean_int64_to_int32(int64_t a) { return (int32_t)a; } +static inline ptrdiff_t lean_int64_to_isize(int64_t a) { return (ptrdiff_t)a; } /* ISize */ LEAN_EXPORT ptrdiff_t lean_isize_of_big_int(b_lean_obj_arg a); -static inline size_t lean_isize_of_int(b_lean_obj_arg a) { +static inline ptrdiff_t lean_isize_of_int(b_lean_obj_arg a) { ptrdiff_t res; if (lean_is_scalar(a)) { @@ -2595,10 +2518,10 @@ static inline size_t lean_isize_of_int(b_lean_obj_arg a) { res = lean_isize_of_big_int(a); } - return (size_t)res; + return res; } -static inline size_t lean_isize_of_nat(b_lean_obj_arg a) { +static inline ptrdiff_t lean_isize_of_nat(b_lean_obj_arg a) { ptrdiff_t res; if (lean_is_scalar(a)) { @@ -2607,132 +2530,100 @@ static inline size_t lean_isize_of_nat(b_lean_obj_arg a) { res = lean_isize_of_big_int(a); } - return (size_t)res; + return res; } -static inline lean_obj_res lean_isize_to_int(size_t a) { - ptrdiff_t arg = (ptrdiff_t)a; - return lean_int64_to_int((int64_t)arg); +static inline ptrdiff_t lean_isize_of_usize_mk(size_t a) { + return (ptrdiff_t)a; } -static inline size_t lean_isize_neg(size_t a) { - ptrdiff_t arg = (ptrdiff_t)a; - - return (size_t)(-arg); +static inline size_t lean_isize_to_usize(ptrdiff_t a) { + return (size_t)a; } -static inline size_t lean_isize_add(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(lhs + rhs); +static inline lean_obj_res lean_isize_to_int(ptrdiff_t a) { + return lean_int64_to_int(a); } -static inline size_t lean_isize_sub(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(lhs - rhs); +static inline ptrdiff_t lean_isize_neg(ptrdiff_t a) { + return -a; } -static inline size_t lean_isize_mul(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(lhs * rhs); +static inline ptrdiff_t lean_isize_add(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs + rhs; } -static inline size_t lean_isize_div(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(rhs == 0 ? 0 : lhs / rhs); +static inline ptrdiff_t lean_isize_sub(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs - rhs; } -static inline size_t lean_isize_mod(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(rhs == 0 ? lhs : lhs % rhs); +static inline ptrdiff_t lean_isize_mul(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs * rhs; } -static inline size_t lean_isize_land(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - - return (size_t)(lhs & rhs); +static inline ptrdiff_t lean_isize_div(ptrdiff_t lhs, ptrdiff_t rhs) { + return rhs == 0 ? 0 : lhs / rhs; } -static inline size_t lean_isize_lor(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; +static inline ptrdiff_t lean_isize_mod(ptrdiff_t lhs, ptrdiff_t rhs) { + return rhs == 0 ? lhs : lhs % rhs; +} - return (size_t)(lhs | rhs); +static inline ptrdiff_t lean_isize_land(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs & rhs; } -static inline size_t lean_isize_xor(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; +static inline ptrdiff_t lean_isize_lor(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs | rhs; +} - return (size_t)(lhs ^ rhs); +static inline ptrdiff_t lean_isize_xor(ptrdiff_t lhs, ptrdiff_t rhs) { + return lhs ^ rhs; } -static inline size_t lean_isize_shift_right(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; +static inline ptrdiff_t lean_isize_shift_right(ptrdiff_t a1, ptrdiff_t a2) { + ptrdiff_t lhs = a1; ptrdiff_t size = sizeof(ptrdiff_t) * 8; - ptrdiff_t rhs = (((ptrdiff_t)a2 % size) + size) % size; // this is smod + ptrdiff_t rhs = ((a2 % size) + size) % size; // this is smod - return (size_t)(lhs >> rhs); + return lhs >> rhs; } -static inline size_t lean_isize_shift_left(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; +static inline ptrdiff_t lean_isize_shift_left(ptrdiff_t a1, ptrdiff_t a2) { + ptrdiff_t lhs = a1; ptrdiff_t size = sizeof(ptrdiff_t) * 8; - ptrdiff_t rhs = (((ptrdiff_t)a2 % size) + size) % size; // this is smod + ptrdiff_t rhs = ((a2 % size) + size) % size; // this is smod - return (size_t)(lhs << rhs); + return lhs << rhs; } -static inline size_t lean_isize_complement(size_t a) { - ptrdiff_t arg = (ptrdiff_t)a; - - return (size_t)(~arg); +static inline ptrdiff_t lean_isize_complement(ptrdiff_t a) { + return ~a; } -static inline size_t lean_isize_abs(size_t a) { - ptrdiff_t arg = (ptrdiff_t)a; - +static inline ptrdiff_t lean_isize_abs(ptrdiff_t a) { // Recall that we are compiling with -fwrapv so this is guaranteed to // map ISIZE_MIN to ISIZE_MIN - return (size_t)(arg < 0 ? -arg : arg); + return a < 0 ? -a : a; } -static inline uint8_t lean_isize_dec_eq(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - +static inline uint8_t lean_isize_dec_eq(ptrdiff_t lhs, ptrdiff_t rhs) { return lhs == rhs; } -static inline uint8_t lean_isize_dec_lt(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - +static inline uint8_t lean_isize_dec_lt(ptrdiff_t lhs, ptrdiff_t rhs) { return lhs < rhs; } -static inline uint8_t lean_isize_dec_le(size_t a1, size_t a2) { - ptrdiff_t lhs = (ptrdiff_t)a1; - ptrdiff_t rhs = (ptrdiff_t)a2; - +static inline uint8_t lean_isize_dec_le(ptrdiff_t lhs, ptrdiff_t rhs) { return lhs <= rhs; } /* ISize -> other */ -static inline uint8_t lean_isize_to_int8(size_t a) { return (uint8_t)(int8_t)(ptrdiff_t)a; } -static inline uint16_t lean_isize_to_int16(size_t a) { return (uint16_t)(int16_t)(ptrdiff_t)a; } -static inline uint32_t lean_isize_to_int32(size_t a) { return (uint32_t)(int32_t)(ptrdiff_t)a; } -static inline uint64_t lean_isize_to_int64(size_t a) { return (uint64_t)(int64_t)(ptrdiff_t)a; } +static inline int8_t lean_isize_to_int8(ptrdiff_t a) { return (int8_t)a; } +static inline int16_t lean_isize_to_int16(ptrdiff_t a) { return (int16_t)a; } +static inline int32_t lean_isize_to_int32(ptrdiff_t a) { return (int32_t)a; } +static inline int64_t lean_isize_to_int64(ptrdiff_t a) { return (int64_t)a; } /* Float */ @@ -2776,6 +2667,28 @@ static inline unsigned lean_unbox_uint32(b_lean_obj_arg o) { } } +static inline lean_obj_res lean_box_int32(int32_t v) { + if (sizeof(void*) == 4) { + /* 32-bit implementation */ + lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(int32_t)); + lean_ctor_set_int32(r, 0, v); + return r; + } else { + /* 64-bit implementation */ + return lean_box(v); + } +} + +static inline unsigned lean_unbox_int32(b_lean_obj_arg o) { + if (sizeof(void*) == 4) { + /* 32-bit implementation */ + return lean_ctor_get_int32(o, 0); + } else { + /* 64-bit implementation */ + return lean_unbox(o); + } +} + static inline lean_obj_res lean_box_uint64(uint64_t v) { lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(uint64_t)); lean_ctor_set_uint64(r, 0, v); @@ -2786,6 +2699,16 @@ static inline uint64_t lean_unbox_uint64(b_lean_obj_arg o) { return lean_ctor_get_uint64(o, 0); } +static inline lean_obj_res lean_box_int64(int64_t v) { + lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(int64_t)); + lean_ctor_set_int64(r, 0, v); + return r; +} + +static inline int64_t lean_unbox_int64(b_lean_obj_arg o) { + return lean_ctor_get_int64(o, 0); +} + static inline lean_obj_res lean_box_usize(size_t v) { lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(size_t)); lean_ctor_set_usize(r, 0, v); @@ -2796,6 +2719,16 @@ static inline size_t lean_unbox_usize(b_lean_obj_arg o) { return lean_ctor_get_usize(o, 0); } +static inline lean_obj_res lean_box_isize(ptrdiff_t v) { + lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(ptrdiff_t)); + lean_ctor_set_isize(r, 0, v); + return r; +} + +static inline ptrdiff_t lean_unbox_isize(b_lean_obj_arg o) { + return lean_ctor_get_isize(o, 0); +} + static inline lean_obj_res lean_box_float(double v) { lean_obj_res r = lean_alloc_ctor(0, 0, sizeof(double)); // NOLINT lean_ctor_set_float(r, 0, v); @@ -2938,41 +2871,41 @@ static inline size_t lean_float_to_usize(double a) { else return (size_t) lean_float_to_uint32(a); // NOLINT } -static inline uint8_t lean_float_to_int8(double a) { +static inline int8_t lean_float_to_int8(double a) { int8_t result; if (lean_float_isnan(a)) result = 0; else result = -129. < a ? (a < 128. ? (int8_t)a : INT8_MAX) : INT8_MIN; - return (uint8_t)result; + return result; } -static inline uint16_t lean_float_to_int16(double a) { +static inline int16_t lean_float_to_int16(double a) { int16_t result; if (lean_float_isnan(a)) result = 0; else result = -32769. < a ? (a < 32768. ? (int16_t)a : INT16_MAX) : INT16_MIN; - return (uint16_t)result; + return result; } -static inline uint32_t lean_float_to_int32(double a) { +static inline int32_t lean_float_to_int32(double a) { int32_t result; if (lean_float_isnan(a)) result = 0; else result = -2147483649. < a ? (a < 2147483648. ? (int32_t)a : INT32_MAX) : INT32_MIN; - return (uint32_t)result; + return result; } -static inline uint64_t lean_float_to_int64(double a) { +static inline int64_t lean_float_to_int64(double a) { int64_t result; if (lean_float_isnan(a)) result = 0; else result = -9223372036854775809. < a ? (a < 9223372036854775808. ? (int64_t)a : INT64_MAX) : INT64_MIN; - return (uint64_t)result; + return result; } -static inline size_t lean_float_to_isize(double a) { +static inline ptrdiff_t lean_float_to_isize(double a) { if (sizeof(size_t) == sizeof(uint64_t)) { ptrdiff_t result; if (lean_float_isnan(a)) result = 0; else result = -9223372036854775809. < a ? (a < 9223372036854775808. ? (ptrdiff_t)a : INT64_MAX) : INT64_MIN; - return (size_t)result; + return result; } else { ptrdiff_t result; if (lean_float_isnan(a)) result = 0; else result = -2147483649. < a ? (a < 2147483648. ? (ptrdiff_t)a : INT32_MAX) : INT32_MIN; - return (size_t)result; + return result; } } LEAN_EXPORT double lean_float_of_bits(uint64_t u); @@ -2990,11 +2923,11 @@ static inline double lean_uint16_to_float(uint16_t a) { return (double) a; } static inline double lean_uint32_to_float(uint32_t a) { return (double) a; } static inline double lean_uint64_to_float(uint64_t a) { return (double) a; } static inline double lean_usize_to_float(size_t a) { return (double) a; } -static inline double lean_int8_to_float(uint8_t a) { return (double)(int8_t) a; } -static inline double lean_int16_to_float(uint16_t a) { return (double)(int16_t) a; } -static inline double lean_int32_to_float(uint32_t a) { return (double)(int32_t) a; } -static inline double lean_int64_to_float(uint64_t a) { return (double)(int64_t) a; } -static inline double lean_isize_to_float(size_t a) { return (double)(ptrdiff_t) a; } +static inline double lean_int8_to_float(int8_t a) { return (double) a; } +static inline double lean_int16_to_float(int16_t a) { return (double) a; } +static inline double lean_int32_to_float(int32_t a) { return (double) a; } +static inline double lean_int64_to_float(int64_t a) { return (double) a; } +static inline double lean_isize_to_float(ptrdiff_t a) { return (double) a; } /* float32 primitives */ static inline uint8_t lean_float32_to_uint8(float a) { @@ -3015,41 +2948,41 @@ static inline size_t lean_float32_to_usize(float a) { else return (size_t) lean_float32_to_uint32(a); // NOLINT } -static inline uint8_t lean_float32_to_int8(float a) { +static inline int8_t lean_float32_to_int8(float a) { int8_t result; if (lean_float32_isnan(a)) result = 0; else result = -129. < a ? (a < 128. ? (int8_t)a : INT8_MAX) : INT8_MIN; - return (uint8_t)result; + return result; } -static inline uint16_t lean_float32_to_int16(float a) { +static inline int16_t lean_float32_to_int16(float a) { int16_t result; if (lean_float32_isnan(a)) result = 0; else result = -32769. < a ? (a < 32768. ? (int16_t)a : INT16_MAX) : INT16_MIN; - return (uint16_t)result; + return result; } -static inline uint32_t lean_float32_to_int32(float a) { +static inline int32_t lean_float32_to_int32(float a) { int32_t result; if (lean_float32_isnan(a)) result = 0; else result = -2147483649. < a ? (a < 2147483648. ? (int32_t)a : INT32_MAX) : INT32_MIN; - return (uint32_t)result; + return result; } -static inline uint64_t lean_float32_to_int64(float a) { +static inline int64_t lean_float32_to_int64(float a) { int64_t result; if (lean_float32_isnan(a)) result = 0; else result = -9223372036854775809. < a ? (a < 9223372036854775808. ? (int64_t)a : INT64_MAX) : INT64_MIN; - return (uint64_t)result; + return result; } -static inline size_t lean_float32_to_isize(float a) { +static inline ptrdiff_t lean_float32_to_isize(float a) { if (sizeof(size_t) == sizeof(uint64_t)) { ptrdiff_t result; if (lean_float32_isnan(a)) result = 0; else result = -9223372036854775809. < a ? (a < 9223372036854775808. ? (ptrdiff_t)a : INT64_MAX) : INT64_MIN; - return (size_t)result; + return result; } else { ptrdiff_t result; if (lean_float32_isnan(a)) result = 0; else result = -2147483649. < a ? (a < 2147483648. ? (ptrdiff_t)a : INT32_MAX) : INT32_MIN; - return (size_t)result; + return result; } } LEAN_EXPORT float lean_float32_of_bits(uint32_t u); @@ -3067,11 +3000,11 @@ static inline float lean_uint16_to_float32(uint16_t a) { return (float) a; } static inline float lean_uint32_to_float32(uint32_t a) { return (float) a; } static inline float lean_uint64_to_float32(uint64_t a) { return (float) a; } static inline float lean_usize_to_float32(size_t a) { return (float) a; } -static inline float lean_int8_to_float32(uint8_t a) { return (float)(int8_t) a; } -static inline float lean_int16_to_float32(uint16_t a) { return (float)(int16_t) a; } -static inline float lean_int32_to_float32(uint32_t a) { return (float)(int32_t) a; } -static inline float lean_int64_to_float32(uint64_t a) { return (float)(int64_t) a; } -static inline float lean_isize_to_float32(size_t a) { return (float)(ptrdiff_t) a; } +static inline float lean_int8_to_float32(int8_t a) { return (float) a; } +static inline float lean_int16_to_float32(int16_t a) { return (float) a; } +static inline float lean_int32_to_float32(int32_t a) { return (float) a; } +static inline float lean_int64_to_float32(int64_t a) { return (float) a; } +static inline float lean_isize_to_float32(ptrdiff_t a) { return (float) a; } static inline float lean_float_to_float32(double a) { return (float)a; } static inline double lean_float32_to_float(float a) { return (double)a; } diff --git a/src/library/ir_interpreter.cpp b/src/library/ir_interpreter.cpp index debc248abced..390c08a3acec 100644 --- a/src/library/ir_interpreter.cpp +++ b/src/library/ir_interpreter.cpp @@ -100,6 +100,7 @@ array_ref const & expr_reuse_args(expr const & e) { lean_assert(expr_tag(e) 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); } +bool expr_uproj_signed(expr const & e) { lean_assert(expr_tag(e) == expr_kind::UProj); return get_bool_field(e.raw(), 2); } 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); } @@ -150,6 +151,7 @@ var_id const & fn_body_uset_target(fn_body const & b) { lean_assert(fn_body_tag( 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); } +bool fn_body_uset_signed(fn_body const & b) { lean_assert(fn_body_tag(b) == fn_body_kind::USet); return get_bool_field(b.raw(), 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); } @@ -260,6 +262,36 @@ union value { v.m_float32 = f; return v; } + + static value from_int8(int8_t i) { + value v; + v.m_num = i; + return v; + } + + static value from_int16(int16_t i) { + value v; + v.m_num = i; + return v; + } + + static value from_int32(int32_t i) { + value v; + v.m_num = i; + return v; + } + + static value from_int64(int64_t i) { + value v; + v.m_num = i; + return v; + } + + static value from_isize(ptrdiff_t i) { + value v; + v.m_num = i; + return v; + } }; object * box_t(value v, type t) { @@ -271,6 +303,11 @@ object * box_t(value v, type t) { case type::UInt32: return box_uint32(v.m_num); case type::UInt64: return box_uint64(v.m_num); case type::USize: return box_size_t(v.m_num); + case type::Int8: return box(v.m_num); + case type::Int16: return box(v.m_num); + case type::Int32: return box_int32(v.m_num); + case type::Int64: return box_int64(v.m_num); + case type::ISize: return box_isize(v.m_num); case type::Object: case type::Tagged: case type::TObject: @@ -293,6 +330,11 @@ value unbox_t(object * o, type t) { case type::UInt32: return unbox_uint32(o); case type::UInt64: return unbox_uint64(o); case type::USize: return unbox_size_t(o); + case type::Int8: return value::from_int8(unbox(o)); + case type::Int16: return value::from_int16(unbox(o)); + case type::Int32: return value::from_int32(unbox_int32(o)); + case type::Int64: return value::from_int64(unbox_int64(o)); + case type::ISize: return value::from_isize(unbox_isize(o)); case type::Irrelevant: case type::Void: case type::Object: @@ -511,7 +553,11 @@ class interpreter { case expr_kind::Proj: // object field access 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()); + if (expr_uproj_signed(e)) { + return value::from_isize(cnstr_get_isize(var(expr_uproj_obj(e)).m_obj, expr_uproj_idx(e).get_small_value())); + } else { + return cnstr_get_usize(var(expr_uproj_obj(e)).m_obj, expr_uproj_idx(e).get_small_value()); + } case expr_kind::SProj: { // other unboxed field access size_t offset = expr_sproj_idx(e).get_small_value() * sizeof(void *) + expr_sproj_offset(e).get_small_value(); @@ -523,7 +569,12 @@ class interpreter { case type::UInt16: return cnstr_get_uint16(o, offset); case type::UInt32: return cnstr_get_uint32(o, offset); case type::UInt64: return cnstr_get_uint64(o, offset); + case type::Int8: return value::from_int8(cnstr_get_int8(o, offset)); + case type::Int16: return value::from_int16(cnstr_get_int16(o, offset)); + case type::Int32: return value::from_int32(cnstr_get_int32(o, offset)); + case type::Int64: return value::from_int64(cnstr_get_int64(o, offset)); case type::USize: + case type::ISize: case type::Irrelevant: case type::Void: case type::Object: @@ -596,6 +647,11 @@ class interpreter { case type::Tagged: case type::TObject: return n.to_obj_arg(); + case type::Int8: + case type::Int16: + case type::Int32: + case type::Int64: + case type::ISize: case type::Irrelevant: case type::Void: break; @@ -701,7 +757,11 @@ class interpreter { case fn_body_kind::USet: { // set USize field of unique reference object * o = var(fn_body_uset_target(b)).m_obj; lean_assert(is_exclusive(o)); - cnstr_set_usize(o, fn_body_uset_idx(b).get_small_value(), var(fn_body_uset_source(b)).m_num); + if (fn_body_uset_signed(b)) { + cnstr_set_isize(o, fn_body_uset_idx(b).get_small_value(), var(fn_body_uset_source(b)).m_num); + } else { + cnstr_set_usize(o, fn_body_uset_idx(b).get_small_value(), var(fn_body_uset_source(b)).m_num); + } b = fn_body_uset_cont(b); break; } @@ -718,7 +778,12 @@ class interpreter { case type::UInt16: cnstr_set_uint16(o, offset, v.m_num); break; case type::UInt32: cnstr_set_uint32(o, offset, v.m_num); break; case type::UInt64: cnstr_set_uint64(o, offset, v.m_num); break; + case type::Int8: cnstr_set_int8(o, offset, v.m_num); break; + case type::Int16: cnstr_set_int16(o, offset, v.m_num); break; + case type::Int32: cnstr_set_int32(o, offset, v.m_num); break; + case type::Int64: cnstr_set_int64(o, offset, v.m_num); break; case type::USize: + case type::ISize: case type::Irrelevant: case type::Void: case type::Object: @@ -900,6 +965,11 @@ class interpreter { case type::UInt32: return *static_cast(e.m_native.m_addr); case type::UInt64: return *static_cast(e.m_native.m_addr); case type::USize: return *static_cast(e.m_native.m_addr); + case type::Int8: return *static_cast(e.m_native.m_addr); + case type::Int16: return *static_cast(e.m_native.m_addr); + case type::Int32: return *static_cast(e.m_native.m_addr); + case type::Int64: return *static_cast(e.m_native.m_addr); + case type::ISize: return *static_cast(e.m_native.m_addr); case type::Object: case type::Tagged: case type::TObject: diff --git a/src/library/ir_types.h b/src/library/ir_types.h index 1cbca9d546f5..747eb59b2ff1 100644 --- a/src/library/ir_types.h +++ b/src/library/ir_types.h @@ -9,18 +9,29 @@ Author: Leonardo de Moura namespace lean { namespace ir { /* -inductive IRType -| float | uint8 | uint16 | uint32 | uint64 | usize -| irrelevant | object | tobject -| float32 -| struct (leanTypeName : Option Name) (types : Array IRType) : IRType -| union (leanTypeName : Name) (types : Array IRType) : IRType -| tagged -| void +inductive IRType where + | float | uint8 | uint16 | uint32 | uint64 | usize + | erased | object | tobject + | float32 + | struct (leanTypeName : Option Name) (types : Array IRType) : IRType + | union (leanTypeName : Name) (types : Array IRType) : IRType + | tagged + | void + | int8 | int16 | int32 | int64 | isize + deriving Inhabited, BEq, Repr Remark: we don't create struct/union types from C++. */ -enum class type { Float, UInt8, UInt16, UInt32, UInt64, USize, Irrelevant, Object, TObject, Float32, Struct, Union, Tagged, Void }; +enum class type { + Float, UInt8, UInt16, UInt32, UInt64, USize, + Irrelevant, Object, TObject, + Float32, + Struct, + Union, + Tagged, + Void, + Int8, Int16, Int32, Int64, ISize +}; typedef nat var_id; typedef nat jp_id; diff --git a/src/runtime/object.h b/src/runtime/object.h index a386eb73a37a..ddf144c031f3 100644 --- a/src/runtime/object.h +++ b/src/runtime/object.h @@ -80,16 +80,26 @@ inline void cnstr_set(u_obj_arg o, unsigned i, obj_arg v) { lean_ctor_set(o, i, inline void cnstr_release(u_obj_arg o, unsigned i) { lean_ctor_release(o, i); } inline usize cnstr_get_usize(b_obj_arg o, unsigned i) { return lean_ctor_get_usize(o, i); } inline void cnstr_set_usize(b_obj_arg o, unsigned i, usize v) { lean_ctor_set_usize(o, i, v); } +inline ptrdiff_t cnstr_get_isize(b_obj_arg o, unsigned i) { return lean_ctor_get_isize(o, i); } +inline void cnstr_set_isize(b_obj_arg o, unsigned i, ptrdiff_t v) { lean_ctor_set_isize(o, i, v); } inline uint8 cnstr_get_uint8(b_obj_arg o, unsigned offset) { return lean_ctor_get_uint8(o, offset); } inline uint16 cnstr_get_uint16(b_obj_arg o, unsigned offset) { return lean_ctor_get_uint16(o, offset); } inline uint32 cnstr_get_uint32(b_obj_arg o, unsigned offset) { return lean_ctor_get_uint32(o, offset); } inline uint64 cnstr_get_uint64(b_obj_arg o, unsigned offset) { return lean_ctor_get_uint64(o, offset); } +inline int8 cnstr_get_int8(b_obj_arg o, unsigned offset) { return lean_ctor_get_int8(o, offset); } +inline int16 cnstr_get_int16(b_obj_arg o, unsigned offset) { return lean_ctor_get_int16(o, offset); } +inline int32 cnstr_get_int32(b_obj_arg o, unsigned offset) { return lean_ctor_get_int32(o, offset); } +inline int64 cnstr_get_int64(b_obj_arg o, unsigned offset) { return lean_ctor_get_int64(o, offset); } inline double cnstr_get_float(b_obj_arg o, unsigned offset) { return lean_ctor_get_float(o, offset); } inline float cnstr_get_float32(b_obj_arg o, unsigned offset) { return lean_ctor_get_float32(o, offset); } inline void cnstr_set_uint8(b_obj_arg o, unsigned offset, uint8 v) { lean_ctor_set_uint8(o, offset, v); } inline void cnstr_set_uint16(b_obj_arg o, unsigned offset, uint16 v) { lean_ctor_set_uint16(o, offset, v); } inline void cnstr_set_uint32(b_obj_arg o, unsigned offset, uint32 v) { lean_ctor_set_uint32(o, offset, v); } inline void cnstr_set_uint64(b_obj_arg o, unsigned offset, uint64 v) { lean_ctor_set_uint64(o, offset, v); } +inline void cnstr_set_int8(b_obj_arg o, unsigned offset, int8 v) { lean_ctor_set_int8(o, offset, v); } +inline void cnstr_set_int16(b_obj_arg o, unsigned offset, int16 v) { lean_ctor_set_int16(o, offset, v); } +inline void cnstr_set_int32(b_obj_arg o, unsigned offset, int32 v) { lean_ctor_set_int32(o, offset, v); } +inline void cnstr_set_int64(b_obj_arg o, unsigned offset, int64 v) { lean_ctor_set_int64(o, offset, v); } inline void cnstr_set_float(b_obj_arg o, unsigned offset, double v) { lean_ctor_set_float(o, offset, v); } inline void cnstr_set_float32(b_obj_arg o, unsigned offset, float v) { lean_ctor_set_float32(o, offset, v); } @@ -375,14 +385,20 @@ inline uint8 int_dec_nonneg(b_obj_arg a) { return lean_int_dec_nonneg(a); } inline obj_res box_uint32(unsigned v) { return lean_box_uint32(v); } inline unsigned unbox_uint32(b_obj_arg o) { return lean_unbox_uint32(o); } +inline obj_res box_int32(int32_t v) { return lean_box_int32(v); } +inline int32_t unbox_int32(b_obj_arg o) { return lean_unbox_int32(o); } inline obj_res box_uint64(unsigned long long v) { return lean_box_uint64(v); } inline unsigned long long unbox_uint64(b_obj_arg o) { return lean_unbox_uint64(o); } +inline obj_res box_int64(int64_t v) { return lean_box_int64(v); } +inline int64_t unbox_int64(b_obj_arg o) { return lean_unbox_int64(o); } inline obj_res box_float(double v) { return lean_box_float(v); } inline obj_res box_float32(float v) { return lean_box_float32(v); } inline double unbox_float(b_obj_arg o) { return lean_unbox_float(o); } inline float unbox_float32(b_obj_arg o) { return lean_unbox_float32(o); } inline obj_res box_size_t(size_t v) { return lean_box_usize(v); } inline size_t unbox_size_t(b_obj_arg o) { return lean_unbox_usize(o); } +inline obj_res box_isize(ptrdiff_t v) { return lean_box_isize(v); } +inline ptrdiff_t unbox_isize(b_obj_arg o) { return lean_unbox_isize(o); } // ======================================= // uint8 diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index a37f4cff9c3d..d1b8a1b637e2 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -2028,14 +2028,14 @@ goto block_72; block_23: { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Array_append___redArg(x_17, x_18); +x_19 = l_Array_append___redArg(x_15, x_18); lean_dec_ref(x_18); -lean_inc(x_14); +lean_inc(x_13); x_20 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_20, 0, x_14); -lean_ctor_set(x_20, 1, x_13); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_12); lean_ctor_set(x_20, 2, x_19); -x_21 = l_Lean_Syntax_node4(x_14, x_16, x_12, x_15, x_11, x_20); +x_21 = l_Lean_Syntax_node4(x_13, x_14, x_16, x_17, x_11, 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); @@ -2132,12 +2132,12 @@ if (lean_obj_tag(x_24) == 0) { lean_object* x_68; x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -x_12 = x_33; -x_13 = x_34; -x_14 = x_30; -x_15 = x_67; -x_16 = x_31; -x_17 = x_35; +x_12 = x_34; +x_13 = x_30; +x_14 = x_31; +x_15 = x_35; +x_16 = x_33; +x_17 = x_67; x_18 = x_68; goto block_23; } @@ -2149,12 +2149,12 @@ lean_inc(x_69); lean_dec_ref(x_24); x_70 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; x_71 = lean_array_push(x_70, x_69); -x_12 = x_33; -x_13 = x_34; -x_14 = x_30; -x_15 = x_67; -x_16 = x_31; -x_17 = x_35; +x_12 = x_34; +x_13 = x_30; +x_14 = x_31; +x_15 = x_35; +x_16 = x_33; +x_17 = x_67; x_18 = x_71; goto block_23; } @@ -6434,32 +6434,32 @@ goto block_713; block_179: { lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -lean_inc_ref(x_32); -x_46 = l_Array_append___redArg(x_32, x_45); +lean_inc_ref(x_33); +x_46 = l_Array_append___redArg(x_33, x_45); lean_dec_ref(x_45); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_31); +lean_ctor_set(x_47, 1, x_41); lean_ctor_set(x_47, 2, x_46); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); x_48 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_48, 0, x_40); -lean_ctor_set(x_48, 1, x_31); -lean_ctor_set(x_48, 2, x_32); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_33); x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__6; -x_50 = l_Lean_Name_mkStr4(x_10, x_11, x_39, x_49); +x_50 = l_Lean_Name_mkStr4(x_10, x_11, x_36, x_49); lean_inc_ref(x_48); lean_inc(x_40); x_51 = l_Lean_Syntax_node1(x_40, x_50, x_48); lean_inc(x_40); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_40); -lean_ctor_set(x_52, 1, x_33); +lean_ctor_set(x_52, 1, x_32); x_53 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__0; -x_54 = l_Lean_Name_mkStr4(x_10, x_11, x_36, x_53); +x_54 = l_Lean_Name_mkStr4(x_10, x_11, x_31, x_53); x_55 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__12; lean_inc(x_40); x_56 = lean_alloc_ctor(2, 2, 0); @@ -6485,26 +6485,26 @@ lean_inc(x_22); lean_inc_ref(x_56); lean_inc(x_40); x_63 = l_Lean_Syntax_node5(x_40, x_54, x_56, x_58, x_60, x_22, x_62); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_64 = l_Lean_Syntax_node1(x_40, x_31, x_63); +x_64 = l_Lean_Syntax_node1(x_40, x_41, x_63); x_65 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__3; lean_inc(x_40); x_66 = l_Lean_Syntax_node1(x_40, x_65, x_24); x_67 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__6; x_69 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__7; -lean_inc(x_34); -lean_inc(x_37); -x_70 = l_Lean_addMacroScope(x_37, x_69, x_34); lean_inc(x_43); +lean_inc(x_34); +x_70 = l_Lean_addMacroScope(x_34, x_69, x_43); +lean_inc(x_39); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_28); -lean_ctor_set(x_71, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_71, 1, x_39); +lean_inc(x_38); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_35); +lean_ctor_set(x_72, 1, x_38); lean_inc(x_40); x_73 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_73, 0, x_40); @@ -6518,18 +6518,18 @@ x_75 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; x_77 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__12; x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__13; +lean_inc(x_43); lean_inc(x_34); -lean_inc(x_37); -x_79 = l_Lean_addMacroScope(x_37, x_78, x_34); +x_79 = l_Lean_addMacroScope(x_34, x_78, x_43); x_80 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__14; -lean_inc(x_43); +lean_inc(x_39); x_81 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_81, 1, x_39); +lean_inc(x_38); x_82 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_35); +lean_ctor_set(x_82, 1, x_38); lean_inc(x_40); x_83 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_83, 0, x_40); @@ -6539,9 +6539,9 @@ lean_ctor_set(x_83, 3, x_82); lean_inc_ref(x_48); lean_inc(x_40); x_84 = l_Lean_Syntax_node2(x_40, x_67, x_83, x_48); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_85 = l_Lean_Syntax_node1(x_40, x_31, x_84); +x_85 = l_Lean_Syntax_node1(x_40, x_41, x_84); lean_inc_ref(x_62); lean_inc_ref(x_56); lean_inc(x_40); @@ -6569,9 +6569,9 @@ lean_inc(x_40); x_95 = l_Lean_Syntax_node1(x_40, x_29, x_94); lean_inc(x_40); x_96 = l_Lean_Syntax_node2(x_40, x_90, x_92, x_95); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_97 = l_Lean_Syntax_node1(x_40, x_31, x_96); +x_97 = l_Lean_Syntax_node1(x_40, x_41, x_96); lean_inc_ref(x_62); lean_inc_ref(x_56); lean_inc(x_40); @@ -6592,18 +6592,18 @@ x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_105 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__24; x_106 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__26; x_107 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__27; +lean_inc(x_43); lean_inc(x_34); -lean_inc(x_37); -x_108 = l_Lean_addMacroScope(x_37, x_107, x_34); +x_108 = l_Lean_addMacroScope(x_34, x_107, x_43); x_109 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__28; -lean_inc(x_43); +lean_inc(x_39); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_110, 1, x_39); +lean_inc(x_38); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_35); +lean_ctor_set(x_111, 1, x_38); lean_inc(x_40); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_40); @@ -6620,18 +6620,18 @@ lean_ctor_set(x_115, 0, x_40); lean_ctor_set(x_115, 1, x_114); x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__31; x_117 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; +lean_inc(x_43); lean_inc(x_34); -lean_inc(x_37); -x_118 = l_Lean_addMacroScope(x_37, x_117, x_34); +x_118 = l_Lean_addMacroScope(x_34, x_117, x_43); x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; -lean_inc(x_43); +lean_inc(x_39); x_120 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_120, 1, x_39); +lean_inc(x_38); x_121 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_35); +lean_ctor_set(x_121, 1, x_38); lean_inc(x_40); x_122 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_122, 0, x_40); @@ -6643,18 +6643,18 @@ lean_inc(x_40); x_123 = l_Lean_Syntax_node2(x_40, x_67, x_122, x_48); x_124 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__35; x_125 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; +lean_inc(x_43); lean_inc(x_34); -lean_inc(x_37); -x_126 = l_Lean_addMacroScope(x_37, x_125, x_34); +x_126 = l_Lean_addMacroScope(x_34, x_125, x_43); x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__37; -lean_inc(x_43); +lean_inc(x_39); x_128 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_128, 1, x_39); +lean_inc(x_38); x_129 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_35); +lean_ctor_set(x_129, 1, x_38); lean_inc(x_40); x_130 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_130, 0, x_40); @@ -6669,9 +6669,9 @@ lean_inc(x_40); x_132 = l_Lean_Syntax_node3(x_40, x_105, x_123, x_115, x_131); lean_inc(x_40); x_133 = l_Lean_Syntax_node3(x_40, x_105, x_113, x_115, x_132); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_134 = l_Lean_Syntax_node1(x_40, x_31, x_133); +x_134 = l_Lean_Syntax_node1(x_40, x_41, x_133); lean_inc_ref(x_62); lean_inc_ref(x_56); lean_inc(x_40); @@ -6692,9 +6692,9 @@ lean_inc(x_40); x_141 = l_Lean_Syntax_node1(x_40, x_29, x_140); lean_inc(x_40); x_142 = l_Lean_Syntax_node1(x_40, x_65, x_141); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_143 = l_Lean_Syntax_node3(x_40, x_31, x_103, x_138, x_142); +x_143 = l_Lean_Syntax_node3(x_40, x_41, x_103, x_138, x_142); lean_inc_ref(x_62); lean_inc_ref(x_56); lean_inc(x_40); @@ -6704,17 +6704,17 @@ lean_inc(x_40); x_145 = l_Lean_Syntax_node2(x_40, x_75, x_144, x_88); x_146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__41; x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__42; +lean_inc(x_43); lean_inc(x_34); -lean_inc(x_37); -x_148 = l_Lean_addMacroScope(x_37, x_147, x_34); +x_148 = l_Lean_addMacroScope(x_34, x_147, x_43); x_149 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__43; x_150 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_43); -lean_inc(x_35); +lean_ctor_set(x_150, 1, x_39); +lean_inc(x_38); x_151 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_35); +lean_ctor_set(x_151, 1, x_38); lean_inc(x_40); x_152 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_152, 0, x_40); @@ -6724,15 +6724,15 @@ lean_ctor_set(x_152, 3, x_151); lean_inc_ref(x_48); lean_inc(x_40); x_153 = l_Lean_Syntax_node2(x_40, x_67, x_152, x_48); -lean_inc(x_31); +lean_inc(x_41); lean_inc(x_40); -x_154 = l_Lean_Syntax_node1(x_40, x_31, x_153); +x_154 = l_Lean_Syntax_node1(x_40, x_41, x_153); lean_inc(x_40); x_155 = l_Lean_Syntax_node3(x_40, x_76, x_56, x_154, x_62); lean_inc(x_40); x_156 = l_Lean_Syntax_node2(x_40, x_75, x_155, x_88); lean_inc(x_40); -x_157 = l_Lean_Syntax_node6(x_40, x_31, x_66, x_74, x_89, x_99, x_145, x_156); +x_157 = l_Lean_Syntax_node6(x_40, x_41, x_66, x_74, x_89, x_99, x_145, x_156); x_158 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__19; lean_inc(x_40); x_159 = lean_alloc_ctor(2, 2, 0); @@ -6740,13 +6740,13 @@ lean_ctor_set(x_159, 0, x_40); lean_ctor_set(x_159, 1, x_158); x_160 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__45; x_161 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; -x_162 = l_Lean_addMacroScope(x_37, x_161, x_34); +x_162 = l_Lean_addMacroScope(x_34, x_161, x_43); lean_inc(x_40); x_163 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_163, 0, x_40); lean_ctor_set(x_163, 1, x_160); lean_ctor_set(x_163, 2, x_162); -lean_ctor_set(x_163, 3, x_35); +lean_ctor_set(x_163, 3, x_38); x_164 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; x_165 = lean_array_push(x_164, x_47); lean_inc_ref(x_48); @@ -6762,308 +6762,308 @@ x_173 = lean_array_push(x_172, x_159); x_174 = lean_array_push(x_173, x_163); x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_40); -lean_ctor_set(x_175, 1, x_44); +lean_ctor_set(x_175, 1, x_35); lean_ctor_set(x_175, 2, x_174); x_176 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_176, 0, x_42); +lean_ctor_set(x_176, 0, x_37); lean_ctor_set(x_176, 1, x_175); x_177 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_177, 0, x_38); +lean_ctor_set(x_177, 0, x_44); lean_ctor_set(x_177, 1, x_176); -x_178 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_177, x_2, x_41); +x_178 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_177, x_2, x_42); x_4 = x_178; goto block_9; } block_319: { lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; -lean_inc_ref(x_187); -x_195 = l_Array_append___redArg(x_187, x_194); +lean_inc_ref(x_185); +x_195 = l_Array_append___redArg(x_185, x_194); lean_dec_ref(x_194); +lean_inc(x_180); lean_inc(x_188); -lean_inc(x_181); x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_181); -lean_ctor_set(x_196, 1, x_188); +lean_ctor_set(x_196, 0, x_188); +lean_ctor_set(x_196, 1, x_180); lean_ctor_set(x_196, 2, x_195); +lean_inc(x_180); lean_inc(x_188); -lean_inc(x_181); x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_181); -lean_ctor_set(x_197, 1, x_188); -lean_ctor_set(x_197, 2, x_187); +lean_ctor_set(x_197, 0, x_188); +lean_ctor_set(x_197, 1, x_180); +lean_ctor_set(x_197, 2, x_185); x_198 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__6; -x_199 = l_Lean_Name_mkStr4(x_10, x_11, x_190, x_198); +x_199 = l_Lean_Name_mkStr4(x_10, x_11, x_187, x_198); lean_inc_ref(x_197); -lean_inc(x_181); -x_200 = l_Lean_Syntax_node1(x_181, x_199, x_197); -lean_inc(x_181); +lean_inc(x_188); +x_200 = l_Lean_Syntax_node1(x_188, x_199, x_197); +lean_inc(x_188); x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_181); -lean_ctor_set(x_201, 1, x_180); +lean_ctor_set(x_201, 0, x_188); +lean_ctor_set(x_201, 1, x_182); x_202 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__0; -x_203 = l_Lean_Name_mkStr4(x_10, x_11, x_182, x_202); +x_203 = l_Lean_Name_mkStr4(x_10, x_11, x_189, x_202); x_204 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__12; -lean_inc(x_181); +lean_inc(x_188); x_205 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_205, 0, x_181); +lean_ctor_set(x_205, 0, x_188); lean_ctor_set(x_205, 1, x_204); x_206 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1; -lean_inc(x_181); +lean_inc(x_188); x_207 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_207, 0, x_181); +lean_ctor_set(x_207, 0, x_188); lean_ctor_set(x_207, 1, x_206); x_208 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__16; -lean_inc(x_181); +lean_inc(x_188); x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_181); +lean_ctor_set(x_209, 0, x_188); lean_ctor_set(x_209, 1, x_208); x_210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__30; -lean_inc(x_181); +lean_inc(x_188); x_211 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_211, 0, x_181); +lean_ctor_set(x_211, 0, x_188); lean_ctor_set(x_211, 1, x_210); lean_inc_ref(x_211); lean_inc(x_22); lean_inc_ref(x_205); -lean_inc(x_181); -x_212 = l_Lean_Syntax_node5(x_181, x_203, x_205, x_207, x_209, x_22, x_211); lean_inc(x_188); -lean_inc(x_181); -x_213 = l_Lean_Syntax_node1(x_181, x_188, x_212); +x_212 = l_Lean_Syntax_node5(x_188, x_203, x_205, x_207, x_209, x_22, x_211); +lean_inc(x_180); +lean_inc(x_188); +x_213 = l_Lean_Syntax_node1(x_188, x_180, x_212); x_214 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__3; -lean_inc(x_181); -x_215 = l_Lean_Syntax_node1(x_181, x_214, x_24); +lean_inc(x_188); +x_215 = l_Lean_Syntax_node1(x_188, x_214, x_24); x_216 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; x_217 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__6; x_218 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__7; +lean_inc(x_192); lean_inc(x_183); -lean_inc(x_189); -x_219 = l_Lean_addMacroScope(x_189, x_218, x_183); -lean_inc(x_193); +x_219 = l_Lean_addMacroScope(x_183, x_218, x_192); +lean_inc(x_186); x_220 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_220, 0, x_28); -lean_ctor_set(x_220, 1, x_193); -lean_inc(x_184); +lean_ctor_set(x_220, 1, x_186); +lean_inc(x_181); x_221 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_184); -lean_inc(x_181); +lean_ctor_set(x_221, 1, x_181); +lean_inc(x_188); x_222 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_222, 0, x_181); +lean_ctor_set(x_222, 0, x_188); lean_ctor_set(x_222, 1, x_217); lean_ctor_set(x_222, 2, x_219); lean_ctor_set(x_222, 3, x_221); lean_inc_ref(x_197); -lean_inc(x_181); -x_223 = l_Lean_Syntax_node2(x_181, x_216, x_222, x_197); +lean_inc(x_188); +x_223 = l_Lean_Syntax_node2(x_188, x_216, x_222, x_197); x_224 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__9; x_225 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; x_226 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__12; x_227 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__13; +lean_inc(x_192); lean_inc(x_183); -lean_inc(x_189); -x_228 = l_Lean_addMacroScope(x_189, x_227, x_183); +x_228 = l_Lean_addMacroScope(x_183, x_227, x_192); x_229 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__14; -lean_inc(x_193); +lean_inc(x_186); x_230 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_193); -lean_inc(x_184); +lean_ctor_set(x_230, 1, x_186); +lean_inc(x_181); x_231 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_231, 0, x_230); -lean_ctor_set(x_231, 1, x_184); -lean_inc(x_181); +lean_ctor_set(x_231, 1, x_181); +lean_inc(x_188); x_232 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_232, 0, x_181); +lean_ctor_set(x_232, 0, x_188); lean_ctor_set(x_232, 1, x_226); lean_ctor_set(x_232, 2, x_228); lean_ctor_set(x_232, 3, x_231); lean_inc_ref(x_197); -lean_inc(x_181); -x_233 = l_Lean_Syntax_node2(x_181, x_216, x_232, x_197); lean_inc(x_188); -lean_inc(x_181); -x_234 = l_Lean_Syntax_node1(x_181, x_188, x_233); +x_233 = l_Lean_Syntax_node2(x_188, x_216, x_232, x_197); +lean_inc(x_180); +lean_inc(x_188); +x_234 = l_Lean_Syntax_node1(x_188, x_180, x_233); lean_inc_ref(x_211); lean_inc_ref(x_205); -lean_inc(x_181); -x_235 = l_Lean_Syntax_node3(x_181, x_225, x_205, x_234, x_211); +lean_inc(x_188); +x_235 = l_Lean_Syntax_node3(x_188, x_225, x_205, x_234, x_211); x_236 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__15; -lean_inc(x_181); +lean_inc(x_188); x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_181); +lean_ctor_set(x_237, 0, x_188); lean_ctor_set(x_237, 1, x_236); lean_inc_ref(x_237); -lean_inc(x_181); -x_238 = l_Lean_Syntax_node2(x_181, x_224, x_235, x_237); +lean_inc(x_188); +x_238 = l_Lean_Syntax_node2(x_188, x_224, x_235, x_237); x_239 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__17; x_240 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__18; -lean_inc(x_181); +lean_inc(x_188); x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_181); +lean_ctor_set(x_241, 0, x_188); lean_ctor_set(x_241, 1, x_240); x_242 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__19; -lean_inc(x_181); +lean_inc(x_188); x_243 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_243, 0, x_181); +lean_ctor_set(x_243, 0, x_188); lean_ctor_set(x_243, 1, x_242); -lean_inc(x_181); -x_244 = l_Lean_Syntax_node1(x_181, x_29, x_243); -lean_inc(x_181); -x_245 = l_Lean_Syntax_node2(x_181, x_239, x_241, x_244); lean_inc(x_188); -lean_inc(x_181); -x_246 = l_Lean_Syntax_node1(x_181, x_188, x_245); +x_244 = l_Lean_Syntax_node1(x_188, x_29, x_243); +lean_inc(x_188); +x_245 = l_Lean_Syntax_node2(x_188, x_239, x_241, x_244); +lean_inc(x_180); +lean_inc(x_188); +x_246 = l_Lean_Syntax_node1(x_188, x_180, x_245); lean_inc_ref(x_211); lean_inc_ref(x_205); -lean_inc(x_181); -x_247 = l_Lean_Syntax_node3(x_181, x_225, x_205, x_246, x_211); +lean_inc(x_188); +x_247 = l_Lean_Syntax_node3(x_188, x_225, x_205, x_246, x_211); lean_inc_ref(x_237); -lean_inc(x_181); -x_248 = l_Lean_Syntax_node2(x_181, x_224, x_247, x_237); +lean_inc(x_188); +x_248 = l_Lean_Syntax_node2(x_188, x_224, x_247, x_237); x_249 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__20; -lean_inc(x_181); +lean_inc(x_188); x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_181); +lean_ctor_set(x_250, 0, x_188); lean_ctor_set(x_250, 1, x_249); -lean_inc(x_181); -x_251 = l_Lean_Syntax_node1(x_181, x_29, x_250); -lean_inc(x_181); -x_252 = l_Lean_Syntax_node1(x_181, x_214, x_251); +lean_inc(x_188); +x_251 = l_Lean_Syntax_node1(x_188, x_29, x_250); +lean_inc(x_188); +x_252 = l_Lean_Syntax_node1(x_188, x_214, x_251); x_253 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__22; x_254 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__24; x_255 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__31; x_256 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; +lean_inc(x_192); lean_inc(x_183); -lean_inc(x_189); -x_257 = l_Lean_addMacroScope(x_189, x_256, x_183); +x_257 = l_Lean_addMacroScope(x_183, x_256, x_192); x_258 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; -lean_inc(x_193); +lean_inc(x_186); x_259 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_193); -lean_inc(x_184); +lean_ctor_set(x_259, 1, x_186); +lean_inc(x_181); x_260 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_260, 0, x_259); -lean_ctor_set(x_260, 1, x_184); -lean_inc(x_181); +lean_ctor_set(x_260, 1, x_181); +lean_inc(x_188); x_261 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_261, 0, x_181); +lean_ctor_set(x_261, 0, x_188); lean_ctor_set(x_261, 1, x_255); lean_ctor_set(x_261, 2, x_257); lean_ctor_set(x_261, 3, x_260); lean_inc_ref(x_197); -lean_inc(x_181); -x_262 = l_Lean_Syntax_node2(x_181, x_216, x_261, x_197); +lean_inc(x_188); +x_262 = l_Lean_Syntax_node2(x_188, x_216, x_261, x_197); x_263 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__29; -lean_inc(x_181); +lean_inc(x_188); x_264 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_264, 0, x_181); +lean_ctor_set(x_264, 0, x_188); lean_ctor_set(x_264, 1, x_263); x_265 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__35; x_266 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; +lean_inc(x_192); lean_inc(x_183); -lean_inc(x_189); -x_267 = l_Lean_addMacroScope(x_189, x_266, x_183); +x_267 = l_Lean_addMacroScope(x_183, x_266, x_192); x_268 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__37; -lean_inc(x_193); +lean_inc(x_186); x_269 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_269, 0, x_268); -lean_ctor_set(x_269, 1, x_193); -lean_inc(x_184); +lean_ctor_set(x_269, 1, x_186); +lean_inc(x_181); x_270 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_184); -lean_inc(x_181); +lean_ctor_set(x_270, 1, x_181); +lean_inc(x_188); x_271 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_271, 0, x_181); +lean_ctor_set(x_271, 0, x_188); lean_ctor_set(x_271, 1, x_265); lean_ctor_set(x_271, 2, x_267); lean_ctor_set(x_271, 3, x_270); lean_inc_ref(x_197); -lean_inc(x_181); -x_272 = l_Lean_Syntax_node2(x_181, x_216, x_271, x_197); -lean_inc(x_181); -x_273 = l_Lean_Syntax_node3(x_181, x_254, x_262, x_264, x_272); lean_inc(x_188); -lean_inc(x_181); -x_274 = l_Lean_Syntax_node1(x_181, x_188, x_273); +x_272 = l_Lean_Syntax_node2(x_188, x_216, x_271, x_197); +lean_inc(x_188); +x_273 = l_Lean_Syntax_node3(x_188, x_254, x_262, x_264, x_272); +lean_inc(x_180); +lean_inc(x_188); +x_274 = l_Lean_Syntax_node1(x_188, x_180, x_273); lean_inc_ref(x_211); lean_inc_ref(x_205); -lean_inc(x_181); -x_275 = l_Lean_Syntax_node3(x_181, x_225, x_205, x_274, x_211); +lean_inc(x_188); +x_275 = l_Lean_Syntax_node3(x_188, x_225, x_205, x_274, x_211); x_276 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__38; -lean_inc(x_181); +lean_inc(x_188); x_277 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_277, 0, x_181); +lean_ctor_set(x_277, 0, x_188); lean_ctor_set(x_277, 1, x_276); -lean_inc(x_181); -x_278 = l_Lean_Syntax_node2(x_181, x_253, x_275, x_277); +lean_inc(x_188); +x_278 = l_Lean_Syntax_node2(x_188, x_253, x_275, x_277); x_279 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__39; -lean_inc(x_181); +lean_inc(x_188); x_280 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_280, 0, x_181); +lean_ctor_set(x_280, 0, x_188); lean_ctor_set(x_280, 1, x_279); -lean_inc(x_181); -x_281 = l_Lean_Syntax_node1(x_181, x_29, x_280); -lean_inc(x_181); -x_282 = l_Lean_Syntax_node1(x_181, x_214, x_281); lean_inc(x_188); -lean_inc(x_181); -x_283 = l_Lean_Syntax_node3(x_181, x_188, x_252, x_278, x_282); +x_281 = l_Lean_Syntax_node1(x_188, x_29, x_280); +lean_inc(x_188); +x_282 = l_Lean_Syntax_node1(x_188, x_214, x_281); +lean_inc(x_180); +lean_inc(x_188); +x_283 = l_Lean_Syntax_node3(x_188, x_180, x_252, x_278, x_282); lean_inc_ref(x_211); lean_inc_ref(x_205); -lean_inc(x_181); -x_284 = l_Lean_Syntax_node3(x_181, x_225, x_205, x_283, x_211); +lean_inc(x_188); +x_284 = l_Lean_Syntax_node3(x_188, x_225, x_205, x_283, x_211); lean_inc_ref(x_237); -lean_inc(x_181); -x_285 = l_Lean_Syntax_node2(x_181, x_224, x_284, x_237); +lean_inc(x_188); +x_285 = l_Lean_Syntax_node2(x_188, x_224, x_284, x_237); x_286 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__41; x_287 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__42; +lean_inc(x_192); lean_inc(x_183); -lean_inc(x_189); -x_288 = l_Lean_addMacroScope(x_189, x_287, x_183); +x_288 = l_Lean_addMacroScope(x_183, x_287, x_192); x_289 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__43; x_290 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_290, 0, x_289); -lean_ctor_set(x_290, 1, x_193); -lean_inc(x_184); +lean_ctor_set(x_290, 1, x_186); +lean_inc(x_181); x_291 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_291, 0, x_290); -lean_ctor_set(x_291, 1, x_184); -lean_inc(x_181); +lean_ctor_set(x_291, 1, x_181); +lean_inc(x_188); x_292 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_292, 0, x_181); +lean_ctor_set(x_292, 0, x_188); lean_ctor_set(x_292, 1, x_286); lean_ctor_set(x_292, 2, x_288); lean_ctor_set(x_292, 3, x_291); lean_inc_ref(x_197); -lean_inc(x_181); -x_293 = l_Lean_Syntax_node2(x_181, x_216, x_292, x_197); lean_inc(x_188); -lean_inc(x_181); -x_294 = l_Lean_Syntax_node1(x_181, x_188, x_293); -lean_inc(x_181); -x_295 = l_Lean_Syntax_node3(x_181, x_225, x_205, x_294, x_211); -lean_inc(x_181); -x_296 = l_Lean_Syntax_node2(x_181, x_224, x_295, x_237); -lean_inc(x_181); -x_297 = l_Lean_Syntax_node6(x_181, x_188, x_215, x_223, x_238, x_248, x_285, x_296); +x_293 = l_Lean_Syntax_node2(x_188, x_216, x_292, x_197); +lean_inc(x_180); +lean_inc(x_188); +x_294 = l_Lean_Syntax_node1(x_188, x_180, x_293); +lean_inc(x_188); +x_295 = l_Lean_Syntax_node3(x_188, x_225, x_205, x_294, x_211); +lean_inc(x_188); +x_296 = l_Lean_Syntax_node2(x_188, x_224, x_295, x_237); +lean_inc(x_188); +x_297 = l_Lean_Syntax_node6(x_188, x_180, x_215, x_223, x_238, x_248, x_285, x_296); x_298 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__19; -lean_inc(x_181); +lean_inc(x_188); x_299 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_299, 0, x_181); +lean_ctor_set(x_299, 0, x_188); lean_ctor_set(x_299, 1, x_298); x_300 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__45; x_301 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; -x_302 = l_Lean_addMacroScope(x_189, x_301, x_183); -lean_inc(x_181); +x_302 = l_Lean_addMacroScope(x_183, x_301, x_192); +lean_inc(x_188); x_303 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_303, 0, x_181); +lean_ctor_set(x_303, 0, x_188); lean_ctor_set(x_303, 1, x_300); lean_ctor_set(x_303, 2, x_302); -lean_ctor_set(x_303, 3, x_184); +lean_ctor_set(x_303, 3, x_181); x_304 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; x_305 = lean_array_push(x_304, x_196); lean_inc_ref(x_197); @@ -7078,276 +7078,276 @@ x_312 = lean_array_push(x_311, x_297); x_313 = lean_array_push(x_312, x_299); x_314 = lean_array_push(x_313, x_303); x_315 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_315, 0, x_181); -lean_ctor_set(x_315, 1, x_186); +lean_ctor_set(x_315, 0, x_188); +lean_ctor_set(x_315, 1, x_190); lean_ctor_set(x_315, 2, x_314); x_316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_316, 0, x_185); +lean_ctor_set(x_316, 0, x_184); lean_ctor_set(x_316, 1, x_315); x_317 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_317, 0, x_191); +lean_ctor_set(x_317, 0, x_193); lean_ctor_set(x_317, 1, x_316); -x_318 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_317, x_2, x_192); +x_318 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_317, x_2, x_191); x_4 = x_318; goto block_9; } block_448: { lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; 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_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; -lean_inc_ref(x_321); -x_335 = l_Array_append___redArg(x_321, x_334); +lean_inc_ref(x_333); +x_335 = l_Array_append___redArg(x_333, x_334); lean_dec_ref(x_334); -lean_inc(x_320); -lean_inc(x_332); +lean_inc(x_331); +lean_inc(x_330); x_336 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_336, 0, x_332); -lean_ctor_set(x_336, 1, x_320); +lean_ctor_set(x_336, 0, x_330); +lean_ctor_set(x_336, 1, x_331); lean_ctor_set(x_336, 2, x_335); -lean_inc(x_320); -lean_inc(x_332); +lean_inc(x_331); +lean_inc(x_330); x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_332); -lean_ctor_set(x_337, 1, x_320); -lean_ctor_set(x_337, 2, x_321); +lean_ctor_set(x_337, 0, x_330); +lean_ctor_set(x_337, 1, x_331); +lean_ctor_set(x_337, 2, x_333); x_338 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__6; -x_339 = l_Lean_Name_mkStr4(x_10, x_11, x_323, x_338); +x_339 = l_Lean_Name_mkStr4(x_10, x_11, x_326, x_338); lean_inc_ref(x_337); -lean_inc(x_332); -x_340 = l_Lean_Syntax_node1(x_332, x_339, x_337); -lean_inc(x_332); +lean_inc(x_330); +x_340 = l_Lean_Syntax_node1(x_330, x_339, x_337); +lean_inc(x_330); x_341 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_341, 0, x_332); -lean_ctor_set(x_341, 1, x_328); +lean_ctor_set(x_341, 0, x_330); +lean_ctor_set(x_341, 1, x_329); x_342 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__0; -x_343 = l_Lean_Name_mkStr4(x_10, x_11, x_327, x_342); +x_343 = l_Lean_Name_mkStr4(x_10, x_11, x_332, x_342); x_344 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__12; -lean_inc(x_332); +lean_inc(x_330); x_345 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_345, 0, x_332); +lean_ctor_set(x_345, 0, x_330); lean_ctor_set(x_345, 1, x_344); x_346 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1; -lean_inc(x_332); +lean_inc(x_330); x_347 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_347, 0, x_332); +lean_ctor_set(x_347, 0, x_330); lean_ctor_set(x_347, 1, x_346); x_348 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__16; -lean_inc(x_332); +lean_inc(x_330); x_349 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_349, 0, x_332); +lean_ctor_set(x_349, 0, x_330); lean_ctor_set(x_349, 1, x_348); x_350 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__30; -lean_inc(x_332); +lean_inc(x_330); x_351 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_351, 0, x_332); +lean_ctor_set(x_351, 0, x_330); lean_ctor_set(x_351, 1, x_350); lean_inc_ref(x_351); lean_inc(x_22); lean_inc_ref(x_345); -lean_inc(x_332); -x_352 = l_Lean_Syntax_node5(x_332, x_343, x_345, x_347, x_349, x_22, x_351); -lean_inc(x_320); -lean_inc(x_332); -x_353 = l_Lean_Syntax_node1(x_332, x_320, x_352); +lean_inc(x_330); +x_352 = l_Lean_Syntax_node5(x_330, x_343, x_345, x_347, x_349, x_22, x_351); +lean_inc(x_331); +lean_inc(x_330); +x_353 = l_Lean_Syntax_node1(x_330, x_331, x_352); x_354 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__3; -lean_inc(x_332); -x_355 = l_Lean_Syntax_node1(x_332, x_354, x_24); +lean_inc(x_330); +x_355 = l_Lean_Syntax_node1(x_330, x_354, x_24); x_356 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; x_357 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__6; x_358 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__7; +lean_inc(x_325); lean_inc(x_324); -lean_inc(x_330); -x_359 = l_Lean_addMacroScope(x_330, x_358, x_324); -lean_inc(x_322); +x_359 = l_Lean_addMacroScope(x_324, x_358, x_325); +lean_inc(x_320); x_360 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_360, 0, x_28); -lean_ctor_set(x_360, 1, x_322); -lean_inc(x_331); +lean_ctor_set(x_360, 1, x_320); +lean_inc(x_323); x_361 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_361, 0, x_360); -lean_ctor_set(x_361, 1, x_331); -lean_inc(x_332); +lean_ctor_set(x_361, 1, x_323); +lean_inc(x_330); x_362 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_362, 0, x_332); +lean_ctor_set(x_362, 0, x_330); lean_ctor_set(x_362, 1, x_357); lean_ctor_set(x_362, 2, x_359); lean_ctor_set(x_362, 3, x_361); lean_inc_ref(x_337); -lean_inc(x_332); -x_363 = l_Lean_Syntax_node2(x_332, x_356, x_362, x_337); +lean_inc(x_330); +x_363 = l_Lean_Syntax_node2(x_330, x_356, x_362, x_337); x_364 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__9; x_365 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; x_366 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__12; x_367 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__13; +lean_inc(x_325); lean_inc(x_324); -lean_inc(x_330); -x_368 = l_Lean_addMacroScope(x_330, x_367, x_324); +x_368 = l_Lean_addMacroScope(x_324, x_367, x_325); x_369 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__14; -lean_inc(x_322); +lean_inc(x_320); x_370 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_370, 0, x_369); -lean_ctor_set(x_370, 1, x_322); -lean_inc(x_331); +lean_ctor_set(x_370, 1, x_320); +lean_inc(x_323); x_371 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_371, 0, x_370); -lean_ctor_set(x_371, 1, x_331); -lean_inc(x_332); +lean_ctor_set(x_371, 1, x_323); +lean_inc(x_330); x_372 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_372, 0, x_332); +lean_ctor_set(x_372, 0, x_330); lean_ctor_set(x_372, 1, x_366); lean_ctor_set(x_372, 2, x_368); lean_ctor_set(x_372, 3, x_371); lean_inc_ref(x_337); -lean_inc(x_332); -x_373 = l_Lean_Syntax_node2(x_332, x_356, x_372, x_337); -lean_inc(x_320); -lean_inc(x_332); -x_374 = l_Lean_Syntax_node1(x_332, x_320, x_373); +lean_inc(x_330); +x_373 = l_Lean_Syntax_node2(x_330, x_356, x_372, x_337); +lean_inc(x_331); +lean_inc(x_330); +x_374 = l_Lean_Syntax_node1(x_330, x_331, x_373); lean_inc_ref(x_351); lean_inc_ref(x_345); -lean_inc(x_332); -x_375 = l_Lean_Syntax_node3(x_332, x_365, x_345, x_374, x_351); +lean_inc(x_330); +x_375 = l_Lean_Syntax_node3(x_330, x_365, x_345, x_374, x_351); x_376 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__15; -lean_inc(x_332); +lean_inc(x_330); x_377 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_377, 0, x_332); +lean_ctor_set(x_377, 0, x_330); lean_ctor_set(x_377, 1, x_376); lean_inc_ref(x_377); -lean_inc(x_332); -x_378 = l_Lean_Syntax_node2(x_332, x_364, x_375, x_377); +lean_inc(x_330); +x_378 = l_Lean_Syntax_node2(x_330, x_364, x_375, x_377); x_379 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__17; x_380 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__18; -lean_inc(x_332); +lean_inc(x_330); x_381 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_381, 0, x_332); +lean_ctor_set(x_381, 0, x_330); lean_ctor_set(x_381, 1, x_380); x_382 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__19; -lean_inc(x_332); +lean_inc(x_330); x_383 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_383, 0, x_332); +lean_ctor_set(x_383, 0, x_330); lean_ctor_set(x_383, 1, x_382); -lean_inc(x_332); -x_384 = l_Lean_Syntax_node1(x_332, x_29, x_383); -lean_inc(x_332); -x_385 = l_Lean_Syntax_node2(x_332, x_379, x_381, x_384); -lean_inc(x_320); -lean_inc(x_332); -x_386 = l_Lean_Syntax_node1(x_332, x_320, x_385); +lean_inc(x_330); +x_384 = l_Lean_Syntax_node1(x_330, x_29, x_383); +lean_inc(x_330); +x_385 = l_Lean_Syntax_node2(x_330, x_379, x_381, x_384); +lean_inc(x_331); +lean_inc(x_330); +x_386 = l_Lean_Syntax_node1(x_330, x_331, x_385); lean_inc_ref(x_351); lean_inc_ref(x_345); -lean_inc(x_332); -x_387 = l_Lean_Syntax_node3(x_332, x_365, x_345, x_386, x_351); +lean_inc(x_330); +x_387 = l_Lean_Syntax_node3(x_330, x_365, x_345, x_386, x_351); lean_inc_ref(x_377); -lean_inc(x_332); -x_388 = l_Lean_Syntax_node2(x_332, x_364, x_387, x_377); +lean_inc(x_330); +x_388 = l_Lean_Syntax_node2(x_330, x_364, x_387, x_377); x_389 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__20; -lean_inc(x_332); +lean_inc(x_330); x_390 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_390, 0, x_332); +lean_ctor_set(x_390, 0, x_330); lean_ctor_set(x_390, 1, x_389); -lean_inc(x_332); -x_391 = l_Lean_Syntax_node1(x_332, x_29, x_390); -lean_inc(x_332); -x_392 = l_Lean_Syntax_node1(x_332, x_354, x_391); +lean_inc(x_330); +x_391 = l_Lean_Syntax_node1(x_330, x_29, x_390); +lean_inc(x_330); +x_392 = l_Lean_Syntax_node1(x_330, x_354, x_391); x_393 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__22; x_394 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__24; x_395 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__31; x_396 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; +lean_inc(x_325); lean_inc(x_324); -lean_inc(x_330); -x_397 = l_Lean_addMacroScope(x_330, x_396, x_324); +x_397 = l_Lean_addMacroScope(x_324, x_396, x_325); x_398 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; -lean_inc(x_322); +lean_inc(x_320); x_399 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_399, 0, x_398); -lean_ctor_set(x_399, 1, x_322); -lean_inc(x_331); +lean_ctor_set(x_399, 1, x_320); +lean_inc(x_323); x_400 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_400, 0, x_399); -lean_ctor_set(x_400, 1, x_331); -lean_inc(x_332); +lean_ctor_set(x_400, 1, x_323); +lean_inc(x_330); x_401 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_401, 0, x_332); +lean_ctor_set(x_401, 0, x_330); lean_ctor_set(x_401, 1, x_395); lean_ctor_set(x_401, 2, x_397); lean_ctor_set(x_401, 3, x_400); lean_inc_ref(x_337); -lean_inc(x_332); -x_402 = l_Lean_Syntax_node2(x_332, x_356, x_401, x_337); +lean_inc(x_330); +x_402 = l_Lean_Syntax_node2(x_330, x_356, x_401, x_337); x_403 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__29; -lean_inc(x_332); +lean_inc(x_330); x_404 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_404, 0, x_332); +lean_ctor_set(x_404, 0, x_330); lean_ctor_set(x_404, 1, x_403); x_405 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__35; x_406 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; +lean_inc(x_325); lean_inc(x_324); -lean_inc(x_330); -x_407 = l_Lean_addMacroScope(x_330, x_406, x_324); +x_407 = l_Lean_addMacroScope(x_324, x_406, x_325); x_408 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__37; x_409 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_409, 0, x_408); -lean_ctor_set(x_409, 1, x_322); -lean_inc(x_331); +lean_ctor_set(x_409, 1, x_320); +lean_inc(x_323); x_410 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_410, 0, x_409); -lean_ctor_set(x_410, 1, x_331); -lean_inc(x_332); +lean_ctor_set(x_410, 1, x_323); +lean_inc(x_330); x_411 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_411, 0, x_332); +lean_ctor_set(x_411, 0, x_330); lean_ctor_set(x_411, 1, x_405); lean_ctor_set(x_411, 2, x_407); lean_ctor_set(x_411, 3, x_410); lean_inc_ref(x_337); -lean_inc(x_332); -x_412 = l_Lean_Syntax_node2(x_332, x_356, x_411, x_337); -lean_inc(x_332); -x_413 = l_Lean_Syntax_node3(x_332, x_394, x_402, x_404, x_412); -lean_inc(x_320); -lean_inc(x_332); -x_414 = l_Lean_Syntax_node1(x_332, x_320, x_413); +lean_inc(x_330); +x_412 = l_Lean_Syntax_node2(x_330, x_356, x_411, x_337); +lean_inc(x_330); +x_413 = l_Lean_Syntax_node3(x_330, x_394, x_402, x_404, x_412); +lean_inc(x_331); +lean_inc(x_330); +x_414 = l_Lean_Syntax_node1(x_330, x_331, x_413); lean_inc_ref(x_351); lean_inc_ref(x_345); -lean_inc(x_332); -x_415 = l_Lean_Syntax_node3(x_332, x_365, x_345, x_414, x_351); +lean_inc(x_330); +x_415 = l_Lean_Syntax_node3(x_330, x_365, x_345, x_414, x_351); x_416 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__38; -lean_inc(x_332); +lean_inc(x_330); x_417 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_417, 0, x_332); +lean_ctor_set(x_417, 0, x_330); lean_ctor_set(x_417, 1, x_416); -lean_inc(x_332); -x_418 = l_Lean_Syntax_node2(x_332, x_393, x_415, x_417); +lean_inc(x_330); +x_418 = l_Lean_Syntax_node2(x_330, x_393, x_415, x_417); x_419 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__39; -lean_inc(x_332); +lean_inc(x_330); x_420 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_420, 0, x_332); +lean_ctor_set(x_420, 0, x_330); lean_ctor_set(x_420, 1, x_419); -lean_inc(x_332); -x_421 = l_Lean_Syntax_node1(x_332, x_29, x_420); -lean_inc(x_332); -x_422 = l_Lean_Syntax_node1(x_332, x_354, x_421); -lean_inc(x_320); -lean_inc(x_332); -x_423 = l_Lean_Syntax_node3(x_332, x_320, x_392, x_418, x_422); -lean_inc(x_332); -x_424 = l_Lean_Syntax_node3(x_332, x_365, x_345, x_423, x_351); -lean_inc(x_332); -x_425 = l_Lean_Syntax_node2(x_332, x_364, x_424, x_377); -lean_inc(x_332); -x_426 = l_Lean_Syntax_node5(x_332, x_320, x_355, x_363, x_378, x_388, x_425); +lean_inc(x_330); +x_421 = l_Lean_Syntax_node1(x_330, x_29, x_420); +lean_inc(x_330); +x_422 = l_Lean_Syntax_node1(x_330, x_354, x_421); +lean_inc(x_331); +lean_inc(x_330); +x_423 = l_Lean_Syntax_node3(x_330, x_331, x_392, x_418, x_422); +lean_inc(x_330); +x_424 = l_Lean_Syntax_node3(x_330, x_365, x_345, x_423, x_351); +lean_inc(x_330); +x_425 = l_Lean_Syntax_node2(x_330, x_364, x_424, x_377); +lean_inc(x_330); +x_426 = l_Lean_Syntax_node5(x_330, x_331, x_355, x_363, x_378, x_388, x_425); x_427 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2___closed__19; -lean_inc(x_332); +lean_inc(x_330); x_428 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_428, 0, x_332); +lean_ctor_set(x_428, 0, x_330); lean_ctor_set(x_428, 1, x_427); x_429 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__45; x_430 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; -x_431 = l_Lean_addMacroScope(x_330, x_430, x_324); -lean_inc(x_332); +x_431 = l_Lean_addMacroScope(x_324, x_430, x_325); +lean_inc(x_330); x_432 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_432, 0, x_332); +lean_ctor_set(x_432, 0, x_330); lean_ctor_set(x_432, 1, x_429); lean_ctor_set(x_432, 2, x_431); -lean_ctor_set(x_432, 3, x_331); +lean_ctor_set(x_432, 3, x_323); x_433 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; x_434 = lean_array_push(x_433, x_336); lean_inc_ref(x_337); @@ -7362,16 +7362,16 @@ x_441 = lean_array_push(x_440, x_426); x_442 = lean_array_push(x_441, x_428); x_443 = lean_array_push(x_442, x_432); x_444 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_444, 0, x_332); -lean_ctor_set(x_444, 1, x_326); +lean_ctor_set(x_444, 0, x_330); +lean_ctor_set(x_444, 1, x_321); lean_ctor_set(x_444, 2, x_443); x_445 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_445, 0, x_333); +lean_ctor_set(x_445, 0, x_327); lean_ctor_set(x_445, 1, x_444); x_446 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_446, 0, x_329); +lean_ctor_set(x_446, 0, x_328); lean_ctor_set(x_446, 1, x_445); -x_447 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_446, x_2, x_325); +x_447 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lam__2(x_10, x_11, x_22, x_27, x_28, x_26, x_30, x_12, x_446, x_2, x_322); x_4 = x_447; goto block_9; } @@ -7446,22 +7446,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_484; x_484 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_455); lean_inc(x_456); -x_180 = x_480; -x_181 = x_478; -x_182 = x_479; -x_183 = x_456; -x_184 = x_472; -x_185 = x_477; -x_186 = x_481; -x_187 = x_483; -x_188 = x_482; -x_189 = x_455; -x_190 = x_466; -x_191 = x_475; -x_192 = x_465; -x_193 = x_472; +lean_inc(x_455); +x_180 = x_482; +x_181 = x_472; +x_182 = x_480; +x_183 = x_455; +x_184 = x_477; +x_185 = x_483; +x_186 = x_472; +x_187 = x_466; +x_188 = x_478; +x_189 = x_479; +x_190 = x_481; +x_191 = x_465; +x_192 = x_456; +x_193 = x_475; x_194 = x_484; goto block_319; } @@ -7472,22 +7472,22 @@ x_485 = lean_ctor_get(x_449, 0); lean_inc(x_485); lean_dec_ref(x_449); x_486 = l_Array_mkArray1___redArg(x_485); -lean_inc(x_455); lean_inc(x_456); -x_180 = x_480; -x_181 = x_478; -x_182 = x_479; -x_183 = x_456; -x_184 = x_472; -x_185 = x_477; -x_186 = x_481; -x_187 = x_483; -x_188 = x_482; -x_189 = x_455; -x_190 = x_466; -x_191 = x_475; -x_192 = x_465; -x_193 = x_472; +lean_inc(x_455); +x_180 = x_482; +x_181 = x_472; +x_182 = x_480; +x_183 = x_455; +x_184 = x_477; +x_185 = x_483; +x_186 = x_472; +x_187 = x_466; +x_188 = x_478; +x_189 = x_479; +x_190 = x_481; +x_191 = x_465; +x_192 = x_456; +x_193 = x_475; x_194 = x_486; goto block_319; } @@ -7538,22 +7538,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_508; x_508 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_455); lean_inc(x_456); -x_180 = x_504; -x_181 = x_502; -x_182 = x_503; -x_183 = x_456; -x_184 = x_496; -x_185 = x_501; -x_186 = x_505; -x_187 = x_507; -x_188 = x_506; -x_189 = x_455; -x_190 = x_489; -x_191 = x_499; -x_192 = x_488; -x_193 = x_496; +lean_inc(x_455); +x_180 = x_506; +x_181 = x_496; +x_182 = x_504; +x_183 = x_455; +x_184 = x_501; +x_185 = x_507; +x_186 = x_496; +x_187 = x_489; +x_188 = x_502; +x_189 = x_503; +x_190 = x_505; +x_191 = x_488; +x_192 = x_456; +x_193 = x_499; x_194 = x_508; goto block_319; } @@ -7564,22 +7564,22 @@ x_509 = lean_ctor_get(x_449, 0); lean_inc(x_509); lean_dec_ref(x_449); x_510 = l_Array_mkArray1___redArg(x_509); -lean_inc(x_455); lean_inc(x_456); -x_180 = x_504; -x_181 = x_502; -x_182 = x_503; -x_183 = x_456; -x_184 = x_496; -x_185 = x_501; -x_186 = x_505; -x_187 = x_507; -x_188 = x_506; -x_189 = x_455; -x_190 = x_489; -x_191 = x_499; -x_192 = x_488; -x_193 = x_496; +lean_inc(x_455); +x_180 = x_506; +x_181 = x_496; +x_182 = x_504; +x_183 = x_455; +x_184 = x_501; +x_185 = x_507; +x_186 = x_496; +x_187 = x_489; +x_188 = x_502; +x_189 = x_503; +x_190 = x_505; +x_191 = x_488; +x_192 = x_456; +x_193 = x_499; x_194 = x_510; goto block_319; } @@ -7649,22 +7649,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_537; x_537 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_455); lean_inc(x_456); -x_180 = x_533; -x_181 = x_531; -x_182 = x_532; -x_183 = x_456; -x_184 = x_524; -x_185 = x_530; -x_186 = x_534; -x_187 = x_536; -x_188 = x_535; -x_189 = x_455; -x_190 = x_517; -x_191 = x_527; -x_192 = x_515; -x_193 = x_524; +lean_inc(x_455); +x_180 = x_535; +x_181 = x_524; +x_182 = x_533; +x_183 = x_455; +x_184 = x_530; +x_185 = x_536; +x_186 = x_524; +x_187 = x_517; +x_188 = x_531; +x_189 = x_532; +x_190 = x_534; +x_191 = x_515; +x_192 = x_456; +x_193 = x_527; x_194 = x_537; goto block_319; } @@ -7675,22 +7675,22 @@ x_538 = lean_ctor_get(x_449, 0); lean_inc(x_538); lean_dec_ref(x_449); x_539 = l_Array_mkArray1___redArg(x_538); -lean_inc(x_455); lean_inc(x_456); -x_180 = x_533; -x_181 = x_531; -x_182 = x_532; -x_183 = x_456; -x_184 = x_524; -x_185 = x_530; -x_186 = x_534; -x_187 = x_536; -x_188 = x_535; -x_189 = x_455; -x_190 = x_517; -x_191 = x_527; -x_192 = x_515; -x_193 = x_524; +lean_inc(x_455); +x_180 = x_535; +x_181 = x_524; +x_182 = x_533; +x_183 = x_455; +x_184 = x_530; +x_185 = x_536; +x_186 = x_524; +x_187 = x_517; +x_188 = x_531; +x_189 = x_532; +x_190 = x_534; +x_191 = x_515; +x_192 = x_456; +x_193 = x_527; x_194 = x_539; goto block_319; } @@ -7754,22 +7754,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_569; x_569 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_540); lean_inc(x_541); -x_320 = x_567; -x_321 = x_568; -x_322 = x_557; -x_323 = x_551; -x_324 = x_541; -x_325 = x_550; -x_326 = x_566; -x_327 = x_564; -x_328 = x_565; -x_329 = x_560; -x_330 = x_540; -x_331 = x_557; -x_332 = x_563; -x_333 = x_562; +lean_inc(x_540); +x_320 = x_557; +x_321 = x_566; +x_322 = x_550; +x_323 = x_557; +x_324 = x_540; +x_325 = x_541; +x_326 = x_551; +x_327 = x_562; +x_328 = x_560; +x_329 = x_565; +x_330 = x_563; +x_331 = x_567; +x_332 = x_564; +x_333 = x_568; x_334 = x_569; goto block_448; } @@ -7780,22 +7780,22 @@ x_570 = lean_ctor_get(x_449, 0); lean_inc(x_570); lean_dec_ref(x_449); x_571 = l_Array_mkArray1___redArg(x_570); -lean_inc(x_540); lean_inc(x_541); -x_320 = x_567; -x_321 = x_568; -x_322 = x_557; -x_323 = x_551; -x_324 = x_541; -x_325 = x_550; -x_326 = x_566; -x_327 = x_564; -x_328 = x_565; -x_329 = x_560; -x_330 = x_540; -x_331 = x_557; -x_332 = x_563; -x_333 = x_562; +lean_inc(x_540); +x_320 = x_557; +x_321 = x_566; +x_322 = x_550; +x_323 = x_557; +x_324 = x_540; +x_325 = x_541; +x_326 = x_551; +x_327 = x_562; +x_328 = x_560; +x_329 = x_565; +x_330 = x_563; +x_331 = x_567; +x_332 = x_564; +x_333 = x_568; x_334 = x_571; goto block_448; } @@ -7846,22 +7846,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_593; x_593 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_540); lean_inc(x_541); -x_320 = x_591; -x_321 = x_592; -x_322 = x_581; -x_323 = x_574; -x_324 = x_541; -x_325 = x_573; -x_326 = x_590; -x_327 = x_588; -x_328 = x_589; -x_329 = x_584; -x_330 = x_540; -x_331 = x_581; -x_332 = x_587; -x_333 = x_586; +lean_inc(x_540); +x_320 = x_581; +x_321 = x_590; +x_322 = x_573; +x_323 = x_581; +x_324 = x_540; +x_325 = x_541; +x_326 = x_574; +x_327 = x_586; +x_328 = x_584; +x_329 = x_589; +x_330 = x_587; +x_331 = x_591; +x_332 = x_588; +x_333 = x_592; x_334 = x_593; goto block_448; } @@ -7872,22 +7872,22 @@ x_594 = lean_ctor_get(x_449, 0); lean_inc(x_594); lean_dec_ref(x_449); x_595 = l_Array_mkArray1___redArg(x_594); -lean_inc(x_540); lean_inc(x_541); -x_320 = x_591; -x_321 = x_592; -x_322 = x_581; -x_323 = x_574; -x_324 = x_541; -x_325 = x_573; -x_326 = x_590; -x_327 = x_588; -x_328 = x_589; -x_329 = x_584; -x_330 = x_540; -x_331 = x_581; -x_332 = x_587; -x_333 = x_586; +lean_inc(x_540); +x_320 = x_581; +x_321 = x_590; +x_322 = x_573; +x_323 = x_581; +x_324 = x_540; +x_325 = x_541; +x_326 = x_574; +x_327 = x_586; +x_328 = x_584; +x_329 = x_589; +x_330 = x_587; +x_331 = x_591; +x_332 = x_588; +x_333 = x_592; x_334 = x_595; goto block_448; } @@ -7957,22 +7957,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_622; x_622 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_540); lean_inc(x_541); -x_320 = x_620; -x_321 = x_621; -x_322 = x_609; -x_323 = x_602; -x_324 = x_541; -x_325 = x_600; -x_326 = x_619; -x_327 = x_617; -x_328 = x_618; -x_329 = x_612; -x_330 = x_540; -x_331 = x_609; -x_332 = x_616; -x_333 = x_615; +lean_inc(x_540); +x_320 = x_609; +x_321 = x_619; +x_322 = x_600; +x_323 = x_609; +x_324 = x_540; +x_325 = x_541; +x_326 = x_602; +x_327 = x_615; +x_328 = x_612; +x_329 = x_618; +x_330 = x_616; +x_331 = x_620; +x_332 = x_617; +x_333 = x_621; x_334 = x_622; goto block_448; } @@ -7983,22 +7983,22 @@ x_623 = lean_ctor_get(x_449, 0); lean_inc(x_623); lean_dec_ref(x_449); x_624 = l_Array_mkArray1___redArg(x_623); -lean_inc(x_540); lean_inc(x_541); -x_320 = x_620; -x_321 = x_621; -x_322 = x_609; -x_323 = x_602; -x_324 = x_541; -x_325 = x_600; -x_326 = x_619; -x_327 = x_617; -x_328 = x_618; -x_329 = x_612; -x_330 = x_540; -x_331 = x_609; -x_332 = x_616; -x_333 = x_615; +lean_inc(x_540); +x_320 = x_609; +x_321 = x_619; +x_322 = x_600; +x_323 = x_609; +x_324 = x_540; +x_325 = x_541; +x_326 = x_602; +x_327 = x_615; +x_328 = x_612; +x_329 = x_618; +x_330 = x_616; +x_331 = x_620; +x_332 = x_617; +x_333 = x_621; x_334 = x_624; goto block_448; } @@ -8065,22 +8065,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_655; x_655 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_625); lean_inc(x_626); -x_31 = x_653; -x_32 = x_654; -x_33 = x_651; -x_34 = x_626; -x_35 = x_642; -x_36 = x_650; -x_37 = x_625; -x_38 = x_645; -x_39 = x_636; +lean_inc(x_625); +x_31 = x_650; +x_32 = x_651; +x_33 = x_654; +x_34 = x_625; +x_35 = x_652; +x_36 = x_636; +x_37 = x_647; +x_38 = x_642; +x_39 = x_642; x_40 = x_649; -x_41 = x_635; -x_42 = x_647; -x_43 = x_642; -x_44 = x_652; +x_41 = x_653; +x_42 = x_635; +x_43 = x_626; +x_44 = x_645; x_45 = x_655; goto block_179; } @@ -8091,22 +8091,22 @@ x_656 = lean_ctor_get(x_449, 0); lean_inc(x_656); lean_dec_ref(x_449); x_657 = l_Array_mkArray1___redArg(x_656); -lean_inc(x_625); lean_inc(x_626); -x_31 = x_653; -x_32 = x_654; -x_33 = x_651; -x_34 = x_626; -x_35 = x_642; -x_36 = x_650; -x_37 = x_625; -x_38 = x_645; -x_39 = x_636; +lean_inc(x_625); +x_31 = x_650; +x_32 = x_651; +x_33 = x_654; +x_34 = x_625; +x_35 = x_652; +x_36 = x_636; +x_37 = x_647; +x_38 = x_642; +x_39 = x_642; x_40 = x_649; -x_41 = x_635; -x_42 = x_647; -x_43 = x_642; -x_44 = x_652; +x_41 = x_653; +x_42 = x_635; +x_43 = x_626; +x_44 = x_645; x_45 = x_657; goto block_179; } @@ -8158,22 +8158,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_680; x_680 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_625); lean_inc(x_626); -x_31 = x_678; -x_32 = x_679; -x_33 = x_676; -x_34 = x_626; -x_35 = x_667; -x_36 = x_675; -x_37 = x_625; -x_38 = x_670; -x_39 = x_660; +lean_inc(x_625); +x_31 = x_675; +x_32 = x_676; +x_33 = x_679; +x_34 = x_625; +x_35 = x_677; +x_36 = x_660; +x_37 = x_672; +x_38 = x_667; +x_39 = x_667; x_40 = x_674; -x_41 = x_659; -x_42 = x_672; -x_43 = x_667; -x_44 = x_677; +x_41 = x_678; +x_42 = x_659; +x_43 = x_626; +x_44 = x_670; x_45 = x_680; goto block_179; } @@ -8184,22 +8184,22 @@ x_681 = lean_ctor_get(x_449, 0); lean_inc(x_681); lean_dec_ref(x_449); x_682 = l_Array_mkArray1___redArg(x_681); -lean_inc(x_625); lean_inc(x_626); -x_31 = x_678; -x_32 = x_679; -x_33 = x_676; -x_34 = x_626; -x_35 = x_667; -x_36 = x_675; -x_37 = x_625; -x_38 = x_670; -x_39 = x_660; +lean_inc(x_625); +x_31 = x_675; +x_32 = x_676; +x_33 = x_679; +x_34 = x_625; +x_35 = x_677; +x_36 = x_660; +x_37 = x_672; +x_38 = x_667; +x_39 = x_667; x_40 = x_674; -x_41 = x_659; -x_42 = x_672; -x_43 = x_667; -x_44 = x_677; +x_41 = x_678; +x_42 = x_659; +x_43 = x_626; +x_44 = x_670; x_45 = x_682; goto block_179; } @@ -8270,22 +8270,22 @@ if (lean_obj_tag(x_449) == 0) { lean_object* x_710; x_710 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__31; -lean_inc(x_625); lean_inc(x_626); -x_31 = x_708; -x_32 = x_709; -x_33 = x_706; -x_34 = x_626; -x_35 = x_696; -x_36 = x_705; -x_37 = x_625; -x_38 = x_699; -x_39 = x_689; +lean_inc(x_625); +x_31 = x_705; +x_32 = x_706; +x_33 = x_709; +x_34 = x_625; +x_35 = x_707; +x_36 = x_689; +x_37 = x_702; +x_38 = x_696; +x_39 = x_696; x_40 = x_704; -x_41 = x_687; -x_42 = x_702; -x_43 = x_696; -x_44 = x_707; +x_41 = x_708; +x_42 = x_687; +x_43 = x_626; +x_44 = x_699; x_45 = x_710; goto block_179; } @@ -8296,22 +8296,22 @@ x_711 = lean_ctor_get(x_449, 0); lean_inc(x_711); lean_dec_ref(x_449); x_712 = l_Array_mkArray1___redArg(x_711); -lean_inc(x_625); lean_inc(x_626); -x_31 = x_708; -x_32 = x_709; -x_33 = x_706; -x_34 = x_626; -x_35 = x_696; -x_36 = x_705; -x_37 = x_625; -x_38 = x_699; -x_39 = x_689; +lean_inc(x_625); +x_31 = x_705; +x_32 = x_706; +x_33 = x_709; +x_34 = x_625; +x_35 = x_707; +x_36 = x_689; +x_37 = x_702; +x_38 = x_696; +x_39 = x_696; x_40 = x_704; -x_41 = x_687; -x_42 = x_702; -x_43 = x_696; -x_44 = x_707; +x_41 = x_708; +x_42 = x_687; +x_43 = x_626; +x_44 = x_699; x_45 = x_712; goto block_179; } diff --git a/stage0/stdlib/Init/NotationExtra.c b/stage0/stdlib/Init/NotationExtra.c index 6f028dda776c..a975e39b25d5 100644 --- a/stage0/stdlib/Init/NotationExtra.c +++ b/stage0/stdlib/Init/NotationExtra.c @@ -14068,7 +14068,7 @@ return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_76; lean_object* x_77; lean_object* x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_145; lean_object* x_146; uint8_t x_147; +lean_object* x_17; 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_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_76; size_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_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_145; lean_object* x_146; uint8_t x_147; x_17 = lean_unsigned_to_nat(1u); x_18 = lean_unsigned_to_nat(3u); x_19 = l_Lean_Syntax_getArg(x_1, x_18); @@ -14123,113 +14123,113 @@ goto block_144; block_75: { 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; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -lean_inc_ref(x_24); -x_34 = l_Array_append___redArg(x_24, x_33); +lean_inc_ref(x_27); +x_34 = l_Array_append___redArg(x_27, x_33); lean_dec_ref(x_33); -lean_inc(x_32); -lean_inc(x_28); +lean_inc(x_25); +lean_inc(x_21); x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_28); -lean_ctor_set(x_35, 1, x_32); +lean_ctor_set(x_35, 0, x_21); +lean_ctor_set(x_35, 1, x_25); lean_ctor_set(x_35, 2, x_34); -lean_inc(x_28); -x_36 = l_Lean_Syntax_node2(x_28, x_23, x_30, x_35); +lean_inc(x_21); +x_36 = l_Lean_Syntax_node2(x_21, x_22, x_31, x_35); x_37 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__1; x_38 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__2; -lean_inc(x_28); +lean_inc(x_21); x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_28); +lean_ctor_set(x_39, 0, x_21); lean_ctor_set(x_39, 1, x_37); -x_40 = lean_array_size(x_20); -lean_inc_ref(x_24); -lean_inc(x_32); -lean_inc(x_28); -x_41 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1_spec__1(x_4, x_5, x_6, x_28, x_32, x_24, x_40, x_29, x_20); +x_40 = lean_array_size(x_26); +lean_inc_ref(x_27); +lean_inc(x_25); +lean_inc(x_21); +x_41 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1_spec__1(x_4, x_5, x_6, x_21, x_25, x_27, x_40, x_20, x_26); x_42 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; x_43 = l_Lean_mkSepArray(x_41, x_42); lean_dec_ref(x_41); -lean_inc_ref(x_24); -x_44 = l_Array_append___redArg(x_24, x_43); +lean_inc_ref(x_27); +x_44 = l_Array_append___redArg(x_27, x_43); lean_dec_ref(x_43); -lean_inc(x_32); -lean_inc(x_28); +lean_inc(x_25); +lean_inc(x_21); x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_32); +lean_ctor_set(x_45, 0, x_21); +lean_ctor_set(x_45, 1, x_25); lean_ctor_set(x_45, 2, x_44); -lean_inc(x_32); -lean_inc(x_28); +lean_inc(x_25); +lean_inc(x_21); x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_28); -lean_ctor_set(x_46, 1, x_32); -lean_ctor_set(x_46, 2, x_24); +lean_ctor_set(x_46, 0, x_21); +lean_ctor_set(x_46, 1, x_25); +lean_ctor_set(x_46, 2, x_27); lean_inc_ref(x_46); -lean_inc(x_28); -x_47 = l_Lean_Syntax_node3(x_28, x_38, x_39, x_45, x_46); -lean_inc(x_32); -lean_inc(x_28); -x_48 = l_Lean_Syntax_node1(x_28, x_32, x_47); +lean_inc(x_21); +x_47 = l_Lean_Syntax_node3(x_21, x_38, x_39, x_45, x_46); +lean_inc(x_25); +lean_inc(x_21); +x_48 = l_Lean_Syntax_node1(x_21, x_25, x_47); x_49 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__4; lean_inc_ref(x_46); -lean_inc(x_28); -x_50 = l_Lean_Syntax_node1(x_28, x_49, x_46); +lean_inc(x_21); +x_50 = l_Lean_Syntax_node1(x_21, x_49, x_46); lean_inc_ref(x_46); -lean_inc(x_28); -x_51 = l_Lean_Syntax_node6(x_28, x_26, x_27, x_19, x_36, x_48, x_46, x_50); -lean_inc(x_28); -x_52 = l_Lean_Syntax_node2(x_28, x_21, x_12, x_51); +lean_inc(x_21); +x_51 = l_Lean_Syntax_node6(x_21, x_28, x_32, x_19, x_36, x_48, x_46, x_50); +lean_inc(x_21); +x_52 = l_Lean_Syntax_node2(x_21, x_24, x_12, x_51); x_53 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__5; x_54 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__6; -lean_inc(x_28); +lean_inc(x_21); x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_28); +lean_ctor_set(x_55, 0, x_21); lean_ctor_set(x_55, 1, x_53); x_56 = l_unexpandListNil___redArg___closed__2; -lean_inc(x_28); +lean_inc(x_21); x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_28); +lean_ctor_set(x_57, 0, x_21); lean_ctor_set(x_57, 1, x_56); x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___closed__11; -lean_inc_ref(x_25); -x_59 = l_Lean_Name_mkStr4(x_4, x_5, x_25, x_58); +lean_inc_ref(x_29); +x_59 = l_Lean_Name_mkStr4(x_4, x_5, x_29, x_58); x_60 = l_Lean_command____Unif__hint________Where___x7c___x2d_u22a2___00__closed__6; -x_61 = l_Lean_Name_mkStr4(x_4, x_5, x_25, x_60); +x_61 = l_Lean_Name_mkStr4(x_4, x_5, x_29, x_60); lean_inc_ref(x_46); -lean_inc(x_28); -x_62 = l_Lean_Syntax_node1(x_28, x_61, x_46); +lean_inc(x_21); +x_62 = l_Lean_Syntax_node1(x_21, x_61, x_46); x_63 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__7; x_64 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__8; -lean_inc(x_28); +lean_inc(x_21); x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_28); +lean_ctor_set(x_65, 0, x_21); lean_ctor_set(x_65, 1, x_63); -lean_inc(x_28); -x_66 = l_Lean_Syntax_node2(x_28, x_64, x_65, x_46); -lean_inc(x_28); -x_67 = l_Lean_Syntax_node2(x_28, x_59, x_62, x_66); -lean_inc(x_32); -lean_inc(x_28); -x_68 = l_Lean_Syntax_node1(x_28, x_32, x_67); +lean_inc(x_21); +x_66 = l_Lean_Syntax_node2(x_21, x_64, x_65, x_46); +lean_inc(x_21); +x_67 = l_Lean_Syntax_node2(x_21, x_59, x_62, x_66); +lean_inc(x_25); +lean_inc(x_21); +x_68 = l_Lean_Syntax_node1(x_21, x_25, x_67); x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___closed__21; -lean_inc(x_28); +lean_inc(x_21); x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_28); +lean_ctor_set(x_70, 0, x_21); lean_ctor_set(x_70, 1, x_69); -lean_inc(x_32); -lean_inc(x_28); -x_71 = l_Lean_Syntax_node1(x_28, x_32, x_22); -lean_inc(x_28); -x_72 = l_Lean_Syntax_node5(x_28, x_54, x_55, x_57, x_68, x_70, x_71); -x_73 = l_Lean_Syntax_node2(x_28, x_32, x_52, x_72); +lean_inc(x_25); +lean_inc(x_21); +x_71 = l_Lean_Syntax_node1(x_21, x_25, x_23); +lean_inc(x_21); +x_72 = l_Lean_Syntax_node5(x_21, x_54, x_55, x_57, x_68, x_70, x_71); +x_73 = l_Lean_Syntax_node2(x_21, x_25, x_52, x_72); x_74 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_31); +lean_ctor_set(x_74, 1, x_30); return x_74; } block_111: { 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; 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; size_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_84 = lean_ctor_get(x_78, 5); +x_84 = lean_ctor_get(x_81, 5); x_85 = 0; x_86 = l_Lean_mkIdentFrom(x_19, x_83, x_85); x_87 = l_Lean_SourceInfo_fromRef(x_84, x_85); @@ -14246,10 +14246,10 @@ lean_inc(x_87); x_94 = l_Lean_Syntax_node1(x_87, x_91, x_93); x_95 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___closed__15; x_96 = l___private_Init_NotationExtra_0__Lean_expandExplicitBindersAux_loop___redArg___closed__13; -x_97 = lean_array_size(x_81); -x_98 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1_spec__4___redArg(x_97, x_79, x_81); +x_97 = lean_array_size(x_78); +x_98 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1_spec__4___redArg(x_97, x_77, x_78); x_99 = lean_array_size(x_98); -x_100 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1_spec__4___redArg(x_99, x_79, x_98); +x_100 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1_spec__4___redArg(x_99, x_77, x_98); x_101 = l_Array_append___redArg(x_96, x_100); lean_dec_ref(x_100); lean_inc(x_87); @@ -14257,35 +14257,35 @@ x_102 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_102, 0, x_87); lean_ctor_set(x_102, 1, x_88); lean_ctor_set(x_102, 2, x_101); -if (lean_obj_tag(x_76) == 0) +if (lean_obj_tag(x_82) == 0) { lean_object* x_103; x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___closed__33; x_20 = x_77; -x_21 = x_89; -x_22 = x_86; -x_23 = x_95; -x_24 = x_96; -x_25 = x_82; -x_26 = x_90; -x_27 = x_94; -x_28 = x_87; -x_29 = x_79; -x_30 = x_102; -x_31 = x_80; -x_32 = x_88; +x_21 = x_87; +x_22 = x_95; +x_23 = x_86; +x_24 = x_89; +x_25 = x_88; +x_26 = x_80; +x_27 = x_96; +x_28 = x_90; +x_29 = x_76; +x_30 = x_79; +x_31 = x_102; +x_32 = x_94; x_33 = x_103; goto block_75; } else { lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_104 = lean_ctor_get(x_76, 0); +x_104 = lean_ctor_get(x_82, 0); lean_inc(x_104); -lean_dec_ref(x_76); +lean_dec_ref(x_82); x_105 = l___private_Init_NotationExtra_0__Lean_expandExplicitBindersAux_loop___redArg___closed__15; -lean_inc_ref(x_82); -x_106 = l_Lean_Name_mkStr4(x_4, x_5, x_82, x_105); +lean_inc_ref(x_76); +x_106 = l_Lean_Name_mkStr4(x_4, x_5, x_76, x_105); x_107 = l___private_Init_NotationExtra_0__Lean_expandExplicitBindersAux_loop___redArg___closed__17; lean_inc(x_87); x_108 = lean_alloc_ctor(2, 2, 0); @@ -14295,18 +14295,18 @@ lean_inc(x_87); x_109 = l_Lean_Syntax_node2(x_87, x_106, x_108, x_104); x_110 = l_Array_mkArray1___redArg(x_109); x_20 = x_77; -x_21 = x_89; -x_22 = x_86; -x_23 = x_95; -x_24 = x_96; -x_25 = x_82; -x_26 = x_90; -x_27 = x_94; -x_28 = x_87; -x_29 = x_79; -x_30 = x_102; -x_31 = x_80; -x_32 = x_88; +x_21 = x_87; +x_22 = x_95; +x_23 = x_86; +x_24 = x_89; +x_25 = x_88; +x_26 = x_80; +x_27 = x_96; +x_28 = x_90; +x_29 = x_76; +x_30 = x_79; +x_31 = x_102; +x_32 = x_94; x_33 = x_110; goto block_75; } @@ -14352,13 +14352,13 @@ if (x_130 == 0) { lean_object* x_131; x_131 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lam__0(x_129); -x_76 = x_114; -x_77 = x_125; -x_78 = x_115; -x_79 = x_121; -x_80 = x_116; -x_81 = x_127; -x_82 = x_126; +x_76 = x_126; +x_77 = x_121; +x_78 = x_127; +x_79 = x_116; +x_80 = x_125; +x_81 = x_115; +x_82 = x_114; x_83 = x_131; goto block_111; } @@ -14374,13 +14374,13 @@ x_134 = lean_ctor_get(x_132, 0); x_135 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lam__0(x_134); lean_ctor_set(x_132, 0, x_135); x_136 = l_Lean_MacroScopesView_review(x_132); -x_76 = x_114; -x_77 = x_125; -x_78 = x_115; -x_79 = x_121; -x_80 = x_116; -x_81 = x_127; -x_82 = x_126; +x_76 = x_126; +x_77 = x_121; +x_78 = x_127; +x_79 = x_116; +x_80 = x_125; +x_81 = x_115; +x_82 = x_114; x_83 = x_136; goto block_111; } @@ -14403,13 +14403,13 @@ lean_ctor_set(x_142, 1, x_138); lean_ctor_set(x_142, 2, x_139); lean_ctor_set(x_142, 3, x_140); x_143 = l_Lean_MacroScopesView_review(x_142); -x_76 = x_114; -x_77 = x_125; -x_78 = x_115; -x_79 = x_121; -x_80 = x_116; -x_81 = x_127; -x_82 = x_126; +x_76 = x_126; +x_77 = x_121; +x_78 = x_127; +x_79 = x_116; +x_80 = x_125; +x_81 = x_115; +x_82 = x_114; x_83 = x_143; goto block_111; } diff --git a/stage0/stdlib/Lake/CLI/Translate/Lean.c b/stage0/stdlib/Lake/CLI/Translate/Lean.c index a7b1914a453b..9ed33a132e99 100644 --- a/stage0/stdlib/Lake/CLI/Translate/Lean.c +++ b/stage0/stdlib/Lake/CLI/Translate/Lean.c @@ -22149,9 +22149,9 @@ goto block_177; block_225: { size_t x_207; size_t x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_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_207 = lean_array_size(x_205); +x_207 = lean_array_size(x_204); x_208 = 0; -x_209 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_207, x_208, x_205); +x_209 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_207, x_208, x_204); x_210 = lean_box(0); x_211 = l_Lean_SourceInfo_fromRef(x_210, x_206); x_212 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -22178,7 +22178,7 @@ x_222 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_222, 0, x_211); lean_ctor_set(x_222, 1, x_221); x_223 = l_Lean_Syntax_node3(x_211, x_212, x_214, x_220, x_222); -x_224 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_203, x_223, x_204); +x_224 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_203, x_223, x_205); x_179 = x_224; goto block_202; } @@ -22202,8 +22202,8 @@ if (x_234 == 0) { lean_dec(x_232); lean_dec_ref(x_231); -x_204 = x_226; -x_205 = x_230; +x_204 = x_230; +x_205 = x_226; x_206 = x_234; goto block_225; } @@ -22214,8 +22214,8 @@ x_235 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_231); if (x_235 == 0) { -x_204 = x_226; -x_205 = x_230; +x_204 = x_230; +x_205 = x_226; x_206 = x_235; goto block_225; } @@ -22282,9 +22282,9 @@ goto block_236; block_284: { size_t x_266; size_t x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_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; -x_266 = lean_array_size(x_264); +x_266 = lean_array_size(x_263); x_267 = 0; -x_268 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_266, x_267, x_264); +x_268 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_266, x_267, x_263); x_269 = lean_box(0); x_270 = l_Lean_SourceInfo_fromRef(x_269, x_265); x_271 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -22311,7 +22311,7 @@ x_281 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_281, 0, x_270); lean_ctor_set(x_281, 1, x_280); x_282 = l_Lean_Syntax_node3(x_270, x_271, x_273, x_279, x_281); -x_283 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_262, x_282, x_263); +x_283 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_262, x_282, x_264); x_238 = x_283; goto block_261; } @@ -22335,8 +22335,8 @@ if (x_293 == 0) { lean_dec(x_291); lean_dec_ref(x_290); -x_263 = x_285; -x_264 = x_289; +x_263 = x_289; +x_264 = x_285; x_265 = x_293; goto block_284; } @@ -22347,8 +22347,8 @@ x_294 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_290); if (x_294 == 0) { -x_263 = x_285; -x_264 = x_289; +x_263 = x_289; +x_264 = x_285; x_265 = x_294; goto block_284; } @@ -22444,9 +22444,9 @@ goto block_295; block_352: { size_t x_334; size_t x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; -x_334 = lean_array_size(x_331); +x_334 = lean_array_size(x_332); x_335 = 0; -x_336 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_334, x_335, x_331); +x_336 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_334, x_335, x_332); x_337 = lean_box(0); x_338 = l_Lean_SourceInfo_fromRef(x_337, x_333); x_339 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -22473,7 +22473,7 @@ x_349 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_349, 0, x_338); lean_ctor_set(x_349, 1, x_348); x_350 = l_Lean_Syntax_node3(x_338, x_339, x_341, x_347, x_349); -x_351 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_330, x_350, x_332); +x_351 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_330, x_350, x_331); x_319 = x_351; goto block_329; } @@ -22497,8 +22497,8 @@ if (x_361 == 0) { lean_dec(x_359); lean_dec_ref(x_358); -x_331 = x_357; -x_332 = x_353; +x_331 = x_353; +x_332 = x_357; x_333 = x_361; goto block_352; } @@ -22509,8 +22509,8 @@ x_362 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_358); if (x_362 == 0) { -x_331 = x_357; -x_332 = x_353; +x_331 = x_353; +x_332 = x_357; x_333 = x_362; goto block_352; } @@ -22982,7 +22982,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lake_CLI_Translate_Lean_0__Lake_LeanLibConfig_mkDeclFields(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_35; lean_object* x_36; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_69; lean_object* x_81; lean_object* x_82; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_121; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_158; lean_object* x_172; lean_object* x_173; lean_object* x_200; lean_object* x_201; lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; lean_object* x_251; lean_object* x_265; lean_object* x_266; lean_object* x_293; lean_object* x_294; lean_object* x_295; uint8_t x_296; lean_object* x_316; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; lean_object* x_353; lean_object* x_367; lean_object* x_368; lean_object* x_369; uint8_t x_370; lean_object* x_390; lean_object* x_404; lean_object* x_405; lean_object* x_432; lean_object* x_433; lean_object* x_449; lean_object* x_450; lean_object* x_451; uint8_t x_455; lean_object* x_456; uint8_t x_457; lean_object* x_461; lean_object* x_474; lean_object* x_475; lean_object* x_476; uint8_t x_477; lean_object* x_497; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_514; uint8_t x_515; uint8_t x_516; lean_object* x_520; lean_object* x_533; lean_object* x_534; lean_object* x_535; uint8_t x_536; lean_object* x_556; lean_object* x_567; lean_object* x_568; lean_object* x_592; lean_object* x_593; lean_object* x_594; uint8_t x_598; lean_object* x_599; uint8_t x_600; lean_object* x_604; lean_object* x_617; lean_object* x_618; lean_object* x_629; lean_object* x_630; lean_object* x_631; uint8_t x_632; lean_object* x_652; lean_object* x_663; lean_object* x_664; lean_object* x_665; uint8_t x_666; lean_object* x_686; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; uint8_t x_701; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_35; lean_object* x_36; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_69; lean_object* x_81; lean_object* x_82; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_121; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_158; lean_object* x_172; lean_object* x_173; lean_object* x_200; lean_object* x_201; lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; lean_object* x_251; lean_object* x_265; lean_object* x_266; lean_object* x_293; lean_object* x_294; lean_object* x_295; uint8_t x_296; lean_object* x_316; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; lean_object* x_353; lean_object* x_367; lean_object* x_368; lean_object* x_369; uint8_t x_370; lean_object* x_390; lean_object* x_404; lean_object* x_405; lean_object* x_432; lean_object* x_433; lean_object* x_449; lean_object* x_450; lean_object* x_451; uint8_t x_455; lean_object* x_456; uint8_t x_457; lean_object* x_461; lean_object* x_474; lean_object* x_475; lean_object* x_476; uint8_t x_477; lean_object* x_497; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_514; uint8_t x_515; uint8_t x_516; lean_object* x_520; lean_object* x_533; lean_object* x_534; lean_object* x_535; uint8_t x_536; lean_object* x_556; lean_object* x_567; lean_object* x_568; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_598; uint8_t x_599; uint8_t x_600; lean_object* x_604; lean_object* x_617; lean_object* x_618; lean_object* x_629; lean_object* x_630; lean_object* x_631; uint8_t x_632; lean_object* x_652; lean_object* x_663; lean_object* x_664; lean_object* x_665; uint8_t x_666; lean_object* x_686; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; uint8_t x_701; x_4 = l_Lake_LeanLibConfig_srcDir___proj(x_1); x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); @@ -24100,9 +24100,9 @@ goto block_460; block_496: { size_t x_478; size_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; -x_478 = lean_array_size(x_476); +x_478 = lean_array_size(x_475); x_479 = 0; -x_480 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_LeanLibConfig_mkDeclFields_spec__0(x_478, x_479, x_476); +x_480 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_LeanLibConfig_mkDeclFields_spec__0(x_478, x_479, x_475); x_481 = lean_box(0); x_482 = l_Lean_SourceInfo_fromRef(x_481, x_477); x_483 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -24129,7 +24129,7 @@ x_493 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_493, 0, x_482); lean_ctor_set(x_493, 1, x_492); x_494 = l_Lean_Syntax_node3(x_482, x_483, x_485, x_491, x_493); -x_495 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_474, x_494, x_475); +x_495 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_474, x_494, x_476); x_461 = x_495; goto block_473; } @@ -24154,8 +24154,8 @@ if (x_505 == 0) { lean_dec(x_503); lean_dec_ref(x_502); -x_475 = x_497; -x_476 = x_501; +x_475 = x_501; +x_476 = x_497; x_477 = x_505; goto block_496; } @@ -24166,8 +24166,8 @@ x_506 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_502); if (x_506 == 0) { -x_475 = x_497; -x_476 = x_501; +x_475 = x_501; +x_476 = x_497; x_477 = x_506; goto block_496; } @@ -24406,11 +24406,11 @@ goto block_591; { if (x_600 == 0) { -if (x_598 == 0) +if (x_599 == 0) { lean_object* x_601; x_601 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_593 = x_599; +x_593 = x_598; x_594 = x_601; goto block_597; } @@ -24418,14 +24418,14 @@ else { lean_object* x_602; x_602 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_593 = x_599; +x_593 = x_598; x_594 = x_602; goto block_597; } } else { -x_568 = x_599; +x_568 = x_598; goto block_591; } } @@ -24457,8 +24457,8 @@ else uint8_t x_612; uint8_t x_613; x_612 = lean_unbox(x_608); x_613 = lean_unbox(x_608); -x_598 = x_612; -x_599 = x_604; +x_598 = x_604; +x_599 = x_612; x_600 = x_613; goto block_603; } @@ -24468,8 +24468,8 @@ else uint8_t x_614; uint8_t x_615; x_614 = lean_unbox(x_608); x_615 = lean_unbox(x_609); -x_598 = x_614; -x_599 = x_604; +x_598 = x_604; +x_599 = x_614; x_600 = x_615; goto block_603; } @@ -24509,9 +24509,9 @@ goto block_616; block_651: { size_t x_633; size_t x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; -x_633 = lean_array_size(x_630); +x_633 = lean_array_size(x_631); x_634 = 0; -x_635 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_LeanLibConfig_mkDeclFields_spec__3(x_633, x_634, x_630); +x_635 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_LeanLibConfig_mkDeclFields_spec__3(x_633, x_634, x_631); x_636 = lean_box(0); x_637 = l_Lean_SourceInfo_fromRef(x_636, x_632); x_638 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -24538,7 +24538,7 @@ x_648 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_648, 0, x_637); lean_ctor_set(x_648, 1, x_647); x_649 = l_Lean_Syntax_node3(x_637, x_638, x_640, x_646, x_648); -x_650 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_629, x_649, x_631); +x_650 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_629, x_649, x_630); x_618 = x_650; goto block_628; } @@ -24563,8 +24563,8 @@ if (x_660 == 0) { lean_dec(x_658); lean_dec_ref(x_657); -x_630 = x_656; -x_631 = x_652; +x_630 = x_652; +x_631 = x_656; x_632 = x_660; goto block_651; } @@ -24575,8 +24575,8 @@ x_661 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_657); if (x_661 == 0) { -x_630 = x_656; -x_631 = x_652; +x_630 = x_652; +x_631 = x_656; x_632 = x_661; goto block_651; } @@ -25080,9 +25080,9 @@ goto block_80; block_120: { size_t x_102; size_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_102 = lean_array_size(x_100); +x_102 = lean_array_size(x_99); x_103 = 0; -x_104 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_102, x_103, x_100); +x_104 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_102, x_103, x_99); x_105 = lean_box(0); x_106 = l_Lean_SourceInfo_fromRef(x_105, x_101); x_107 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -25109,7 +25109,7 @@ x_117 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_117, 0, x_106); lean_ctor_set(x_117, 1, x_116); x_118 = l_Lean_Syntax_node3(x_106, x_107, x_109, x_115, x_117); -x_119 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_98, x_118, x_99); +x_119 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_98, x_118, x_100); x_82 = x_119; goto block_97; } @@ -25138,8 +25138,8 @@ if (x_132 == 0) { lean_dec(x_130); lean_dec_ref(x_129); -x_99 = x_121; -x_100 = x_128; +x_99 = x_128; +x_100 = x_121; x_101 = x_132; goto block_120; } @@ -25150,8 +25150,8 @@ x_133 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_129); if (x_133 == 0) { -x_99 = x_121; -x_100 = x_128; +x_99 = x_128; +x_100 = x_121; x_101 = x_133; goto block_120; } @@ -25366,9 +25366,9 @@ goto block_199; block_250: { size_t x_232; size_t x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_232 = lean_array_size(x_230); +x_232 = lean_array_size(x_229); x_233 = 0; -x_234 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_232, x_233, x_230); +x_234 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_232, x_233, x_229); x_235 = lean_box(0); x_236 = l_Lean_SourceInfo_fromRef(x_235, x_231); x_237 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -25395,7 +25395,7 @@ x_247 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_247, 0, x_236); lean_ctor_set(x_247, 1, x_246); x_248 = l_Lean_Syntax_node3(x_236, x_237, x_239, x_245, x_247); -x_249 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_228, x_248, x_229); +x_249 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_228, x_248, x_230); x_201 = x_249; goto block_227; } @@ -25424,8 +25424,8 @@ if (x_262 == 0) { lean_dec(x_260); lean_dec_ref(x_259); -x_229 = x_251; -x_230 = x_258; +x_229 = x_258; +x_230 = x_251; x_231 = x_262; goto block_250; } @@ -25436,8 +25436,8 @@ x_263 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_259); if (x_263 == 0) { -x_229 = x_251; -x_230 = x_258; +x_229 = x_258; +x_230 = x_251; x_231 = x_263; goto block_250; } @@ -25509,9 +25509,9 @@ goto block_264; block_315: { size_t x_297; size_t x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; -x_297 = lean_array_size(x_295); +x_297 = lean_array_size(x_294); x_298 = 0; -x_299 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_297, x_298, x_295); +x_299 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_297, x_298, x_294); x_300 = lean_box(0); x_301 = l_Lean_SourceInfo_fromRef(x_300, x_296); x_302 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -25538,7 +25538,7 @@ x_312 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_312, 0, x_301); lean_ctor_set(x_312, 1, x_311); x_313 = l_Lean_Syntax_node3(x_301, x_302, x_304, x_310, x_312); -x_314 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_293, x_313, x_294); +x_314 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_293, x_313, x_295); x_266 = x_314; goto block_292; } @@ -25567,8 +25567,8 @@ if (x_327 == 0) { lean_dec(x_325); lean_dec_ref(x_324); -x_294 = x_316; -x_295 = x_323; +x_294 = x_323; +x_295 = x_316; x_296 = x_327; goto block_315; } @@ -25579,8 +25579,8 @@ x_328 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_324); if (x_328 == 0) { -x_294 = x_316; -x_295 = x_323; +x_294 = x_323; +x_295 = x_316; x_296 = x_328; goto block_315; } @@ -25681,9 +25681,9 @@ goto block_329; block_389: { size_t x_371; size_t x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; -x_371 = lean_array_size(x_368); +x_371 = lean_array_size(x_369); x_372 = 0; -x_373 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_371, x_372, x_368); +x_373 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_371, x_372, x_369); x_374 = lean_box(0); x_375 = l_Lean_SourceInfo_fromRef(x_374, x_370); x_376 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -25710,7 +25710,7 @@ x_386 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_386, 0, x_375); lean_ctor_set(x_386, 1, x_385); x_387 = l_Lean_Syntax_node3(x_375, x_376, x_378, x_384, x_386); -x_388 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_367, x_387, x_369); +x_388 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_367, x_387, x_368); x_353 = x_388; goto block_366; } @@ -25739,8 +25739,8 @@ if (x_401 == 0) { lean_dec(x_399); lean_dec_ref(x_398); -x_368 = x_397; -x_369 = x_390; +x_368 = x_390; +x_369 = x_397; x_370 = x_401; goto block_389; } @@ -25751,8 +25751,8 @@ x_402 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_398); if (x_402 == 0) { -x_368 = x_397; -x_369 = x_390; +x_368 = x_390; +x_369 = x_397; x_370 = x_402; goto block_389; } @@ -27612,7 +27612,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lake_CLI_Translate_Lean_0__Lake_PackageConfig_mkDeclFields(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_36; lean_object* x_37; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_70; lean_object* x_82; lean_object* x_83; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_122; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_159; lean_object* x_173; lean_object* x_174; lean_object* x_201; lean_object* x_202; lean_object* x_229; lean_object* x_230; lean_object* x_231; uint8_t x_232; lean_object* x_252; lean_object* x_266; lean_object* x_267; lean_object* x_294; lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_317; lean_object* x_331; lean_object* x_332; lean_object* x_333; uint8_t x_334; lean_object* x_354; lean_object* x_368; lean_object* x_369; lean_object* x_370; uint8_t x_371; lean_object* x_391; lean_object* x_405; lean_object* x_406; lean_object* x_433; lean_object* x_434; lean_object* x_450; lean_object* x_451; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_473; uint8_t x_474; uint8_t x_475; lean_object* x_479; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_498; uint8_t x_499; uint8_t x_500; lean_object* x_504; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_523; lean_object* x_524; uint8_t x_525; lean_object* x_529; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_548; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_563; uint8_t x_564; uint8_t x_565; lean_object* x_569; lean_object* x_582; lean_object* x_583; lean_object* x_596; lean_object* x_597; lean_object* x_598; uint8_t x_599; lean_object* x_619; lean_object* x_630; lean_object* x_631; lean_object* x_642; lean_object* x_643; lean_object* x_654; lean_object* x_655; lean_object* x_656; uint8_t x_657; lean_object* x_677; lean_object* x_688; lean_object* x_689; lean_object* x_700; lean_object* x_701; lean_object* x_709; lean_object* x_710; lean_object* x_720; lean_object* x_721; lean_object* x_722; uint8_t x_723; lean_object* x_743; lean_object* x_754; lean_object* x_755; lean_object* x_766; lean_object* x_767; lean_object* x_768; uint8_t x_769; lean_object* x_789; lean_object* x_800; lean_object* x_801; lean_object* x_812; lean_object* x_813; lean_object* x_814; uint8_t x_818; lean_object* x_819; uint8_t x_820; lean_object* x_824; lean_object* x_837; lean_object* x_838; lean_object* x_847; lean_object* x_848; lean_object* x_857; lean_object* x_858; lean_object* x_871; lean_object* x_872; lean_object* x_885; lean_object* x_886; lean_object* x_899; lean_object* x_900; lean_object* x_913; lean_object* x_914; lean_object* x_927; lean_object* x_928; lean_object* x_941; lean_object* x_942; lean_object* x_943; uint8_t x_944; lean_object* x_964; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_981; uint8_t x_982; uint8_t x_983; lean_object* x_987; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; uint8_t x_1003; lean_object* x_1023; lean_object* x_1034; lean_object* x_1035; lean_object* x_1044; lean_object* x_1045; lean_object* x_1049; uint8_t x_1050; lean_object* x_1055; uint8_t x_1056; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_36; lean_object* x_37; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_70; lean_object* x_82; lean_object* x_83; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_122; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_159; lean_object* x_173; lean_object* x_174; lean_object* x_201; lean_object* x_202; lean_object* x_229; lean_object* x_230; lean_object* x_231; uint8_t x_232; lean_object* x_252; lean_object* x_266; lean_object* x_267; lean_object* x_294; lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_317; lean_object* x_331; lean_object* x_332; lean_object* x_333; uint8_t x_334; lean_object* x_354; lean_object* x_368; lean_object* x_369; lean_object* x_370; uint8_t x_371; lean_object* x_391; lean_object* x_405; lean_object* x_406; lean_object* x_433; lean_object* x_434; lean_object* x_450; lean_object* x_451; lean_object* x_467; lean_object* x_468; lean_object* x_469; uint8_t x_473; lean_object* x_474; uint8_t x_475; lean_object* x_479; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_498; lean_object* x_499; uint8_t x_500; lean_object* x_504; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_523; uint8_t x_524; uint8_t x_525; lean_object* x_529; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_548; lean_object* x_557; lean_object* x_558; lean_object* x_559; uint8_t x_563; lean_object* x_564; uint8_t x_565; lean_object* x_569; lean_object* x_582; lean_object* x_583; lean_object* x_596; lean_object* x_597; lean_object* x_598; uint8_t x_599; lean_object* x_619; lean_object* x_630; lean_object* x_631; lean_object* x_642; lean_object* x_643; lean_object* x_654; lean_object* x_655; lean_object* x_656; uint8_t x_657; lean_object* x_677; lean_object* x_688; lean_object* x_689; lean_object* x_700; lean_object* x_701; lean_object* x_709; lean_object* x_710; lean_object* x_720; lean_object* x_721; lean_object* x_722; uint8_t x_723; lean_object* x_743; lean_object* x_754; lean_object* x_755; lean_object* x_766; lean_object* x_767; lean_object* x_768; uint8_t x_769; lean_object* x_789; lean_object* x_800; lean_object* x_801; lean_object* x_812; lean_object* x_813; lean_object* x_814; uint8_t x_818; lean_object* x_819; uint8_t x_820; lean_object* x_824; lean_object* x_837; lean_object* x_838; lean_object* x_847; lean_object* x_848; lean_object* x_857; lean_object* x_858; lean_object* x_871; lean_object* x_872; lean_object* x_885; lean_object* x_886; lean_object* x_899; lean_object* x_900; lean_object* x_913; lean_object* x_914; lean_object* x_927; lean_object* x_928; lean_object* x_941; lean_object* x_942; lean_object* x_943; uint8_t x_944; lean_object* x_964; lean_object* x_975; lean_object* x_976; lean_object* x_977; uint8_t x_981; lean_object* x_982; uint8_t x_983; lean_object* x_987; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; uint8_t x_1003; lean_object* x_1023; lean_object* x_1034; lean_object* x_1035; lean_object* x_1044; lean_object* x_1045; lean_object* x_1049; uint8_t x_1050; lean_object* x_1055; uint8_t x_1056; x_5 = l_Lake_PackageConfig_bootstrap___proj(x_1, x_2); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -27982,9 +27982,9 @@ goto block_98; block_158: { size_t x_140; size_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_140 = lean_array_size(x_137); +x_140 = lean_array_size(x_138); x_141 = 0; -x_142 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_140, x_141, x_137); +x_142 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_140, x_141, x_138); x_143 = lean_box(0); x_144 = l_Lean_SourceInfo_fromRef(x_143, x_139); x_145 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -28011,7 +28011,7 @@ x_155 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_155, 0, x_144); lean_ctor_set(x_155, 1, x_154); x_156 = l_Lean_Syntax_node3(x_144, x_145, x_147, x_153, x_155); -x_157 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_136, x_156, x_138); +x_157 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_136, x_156, x_137); x_122 = x_157; goto block_135; } @@ -28040,8 +28040,8 @@ if (x_170 == 0) { lean_dec(x_168); lean_dec_ref(x_167); -x_137 = x_166; -x_138 = x_159; +x_137 = x_159; +x_138 = x_166; x_139 = x_170; goto block_158; } @@ -28052,8 +28052,8 @@ x_171 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_167); if (x_171 == 0) { -x_137 = x_166; -x_138 = x_159; +x_137 = x_159; +x_138 = x_166; x_139 = x_171; goto block_158; } @@ -28411,9 +28411,9 @@ goto block_293; block_353: { size_t x_335; size_t x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_335 = lean_array_size(x_332); +x_335 = lean_array_size(x_333); x_336 = 0; -x_337 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_335, x_336, x_332); +x_337 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_335, x_336, x_333); x_338 = lean_box(0); x_339 = l_Lean_SourceInfo_fromRef(x_338, x_334); x_340 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -28440,7 +28440,7 @@ x_350 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_350, 0, x_339); lean_ctor_set(x_350, 1, x_349); x_351 = l_Lean_Syntax_node3(x_339, x_340, x_342, x_348, x_350); -x_352 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_331, x_351, x_333); +x_352 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_331, x_351, x_332); x_317 = x_352; goto block_330; } @@ -28469,8 +28469,8 @@ if (x_365 == 0) { lean_dec(x_363); lean_dec_ref(x_362); -x_332 = x_361; -x_333 = x_354; +x_332 = x_354; +x_333 = x_361; x_334 = x_365; goto block_353; } @@ -28481,8 +28481,8 @@ x_366 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_362); if (x_366 == 0) { -x_332 = x_361; -x_333 = x_354; +x_332 = x_354; +x_333 = x_361; x_334 = x_366; goto block_353; } @@ -28497,9 +28497,9 @@ goto block_330; block_390: { size_t x_372; size_t x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; -x_372 = lean_array_size(x_369); +x_372 = lean_array_size(x_370); x_373 = 0; -x_374 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_372, x_373, x_369); +x_374 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_372, x_373, x_370); x_375 = lean_box(0); x_376 = l_Lean_SourceInfo_fromRef(x_375, x_371); x_377 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -28526,7 +28526,7 @@ x_387 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_387, 0, x_376); lean_ctor_set(x_387, 1, x_386); x_388 = l_Lean_Syntax_node3(x_376, x_377, x_379, x_385, x_387); -x_389 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_368, x_388, x_370); +x_389 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_368, x_388, x_369); x_354 = x_389; goto block_367; } @@ -28555,8 +28555,8 @@ if (x_402 == 0) { lean_dec(x_400); lean_dec_ref(x_399); -x_369 = x_398; -x_370 = x_391; +x_369 = x_391; +x_370 = x_398; x_371 = x_402; goto block_390; } @@ -28567,8 +28567,8 @@ x_403 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_399); if (x_403 == 0) { -x_369 = x_398; -x_370 = x_391; +x_369 = x_391; +x_370 = x_398; x_371 = x_403; goto block_390; } @@ -28723,11 +28723,11 @@ goto block_466; { if (x_475 == 0) { -if (x_474 == 0) +if (x_473 == 0) { lean_object* x_476; x_476 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_468 = x_473; +x_468 = x_474; x_469 = x_476; goto block_472; } @@ -28735,14 +28735,14 @@ else { lean_object* x_477; x_477 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_468 = x_473; +x_468 = x_474; x_469 = x_477; goto block_472; } } else { -x_451 = x_473; +x_451 = x_474; goto block_466; } } @@ -28774,8 +28774,8 @@ else uint8_t x_487; uint8_t x_488; x_487 = lean_unbox(x_483); x_488 = lean_unbox(x_483); -x_473 = x_479; -x_474 = x_487; +x_473 = x_487; +x_474 = x_479; x_475 = x_488; goto block_478; } @@ -28785,8 +28785,8 @@ else uint8_t x_489; uint8_t x_490; x_489 = lean_unbox(x_483); x_490 = lean_unbox(x_484); -x_473 = x_479; -x_474 = x_489; +x_473 = x_489; +x_474 = x_479; x_475 = x_490; goto block_478; } @@ -28803,11 +28803,11 @@ goto block_491; { if (x_500 == 0) { -if (x_499 == 0) +if (x_498 == 0) { lean_object* x_501; x_501 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_493 = x_498; +x_493 = x_499; x_494 = x_501; goto block_497; } @@ -28815,14 +28815,14 @@ else { lean_object* x_502; x_502 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_493 = x_498; +x_493 = x_499; x_494 = x_502; goto block_497; } } else { -x_479 = x_498; +x_479 = x_499; goto block_491; } } @@ -28854,8 +28854,8 @@ else uint8_t x_512; uint8_t x_513; x_512 = lean_unbox(x_508); x_513 = lean_unbox(x_508); -x_498 = x_504; -x_499 = x_512; +x_498 = x_512; +x_499 = x_504; x_500 = x_513; goto block_503; } @@ -28865,8 +28865,8 @@ else uint8_t x_514; uint8_t x_515; x_514 = lean_unbox(x_508); x_515 = lean_unbox(x_509); -x_498 = x_504; -x_499 = x_514; +x_498 = x_514; +x_499 = x_504; x_500 = x_515; goto block_503; } @@ -28883,11 +28883,11 @@ goto block_516; { if (x_525 == 0) { -if (x_523 == 0) +if (x_524 == 0) { lean_object* x_526; x_526 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_518 = x_524; +x_518 = x_523; x_519 = x_526; goto block_522; } @@ -28895,14 +28895,14 @@ else { lean_object* x_527; x_527 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_518 = x_524; +x_518 = x_523; x_519 = x_527; goto block_522; } } else { -x_504 = x_524; +x_504 = x_523; goto block_516; } } @@ -28934,8 +28934,8 @@ else uint8_t x_537; uint8_t x_538; x_537 = lean_unbox(x_533); x_538 = lean_unbox(x_533); -x_523 = x_537; -x_524 = x_529; +x_523 = x_529; +x_524 = x_537; x_525 = x_538; goto block_528; } @@ -28945,8 +28945,8 @@ else uint8_t x_539; uint8_t x_540; x_539 = lean_unbox(x_533); x_540 = lean_unbox(x_534); -x_523 = x_539; -x_524 = x_529; +x_523 = x_529; +x_524 = x_539; x_525 = x_540; goto block_528; } @@ -29011,11 +29011,11 @@ goto block_556; { if (x_565 == 0) { -if (x_564 == 0) +if (x_563 == 0) { lean_object* x_566; x_566 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_558 = x_563; +x_558 = x_564; x_559 = x_566; goto block_562; } @@ -29023,14 +29023,14 @@ else { lean_object* x_567; x_567 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_558 = x_563; +x_558 = x_564; x_559 = x_567; goto block_562; } } else { -x_548 = x_563; +x_548 = x_564; goto block_556; } } @@ -29062,8 +29062,8 @@ else uint8_t x_577; uint8_t x_578; x_577 = lean_unbox(x_573); x_578 = lean_unbox(x_573); -x_563 = x_569; -x_564 = x_577; +x_563 = x_577; +x_564 = x_569; x_565 = x_578; goto block_568; } @@ -29073,8 +29073,8 @@ else uint8_t x_579; uint8_t x_580; x_579 = lean_unbox(x_573); x_580 = lean_unbox(x_574); -x_563 = x_569; -x_564 = x_579; +x_563 = x_579; +x_564 = x_569; x_565 = x_580; goto block_568; } @@ -29543,9 +29543,9 @@ goto block_753; block_788: { size_t x_770; size_t x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; -x_770 = lean_array_size(x_767); +x_770 = lean_array_size(x_768); x_771 = 0; -x_772 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_770, x_771, x_767); +x_772 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_770, x_771, x_768); x_773 = lean_box(0); x_774 = l_Lean_SourceInfo_fromRef(x_773, x_769); x_775 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -29572,7 +29572,7 @@ x_785 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_785, 0, x_774); lean_ctor_set(x_785, 1, x_784); x_786 = l_Lean_Syntax_node3(x_774, x_775, x_777, x_783, x_785); -x_787 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_766, x_786, x_768); +x_787 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_766, x_786, x_767); x_755 = x_787; goto block_765; } @@ -29597,8 +29597,8 @@ if (x_797 == 0) { lean_dec(x_795); lean_dec_ref(x_794); -x_767 = x_793; -x_768 = x_789; +x_767 = x_789; +x_768 = x_793; x_769 = x_797; goto block_788; } @@ -29609,8 +29609,8 @@ x_798 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_794); if (x_798 == 0) { -x_767 = x_793; -x_768 = x_789; +x_767 = x_789; +x_768 = x_793; x_769 = x_798; goto block_788; } @@ -30000,9 +30000,9 @@ goto block_926; block_963: { size_t x_945; size_t x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; -x_945 = lean_array_size(x_943); +x_945 = lean_array_size(x_942); x_946 = 0; -x_947 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_945, x_946, x_943); +x_947 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Lean_0__Lake_StrPatDescr_toLean_spec__0(x_945, x_946, x_942); x_948 = lean_box(0); x_949 = l_Lean_SourceInfo_fromRef(x_948, x_944); x_950 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanArrayMkStr1___redArg___lam__1___closed__12; @@ -30029,7 +30029,7 @@ x_960 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_960, 0, x_949); lean_ctor_set(x_960, 1, x_959); x_961 = l_Lean_Syntax_node3(x_949, x_950, x_952, x_958, x_960); -x_962 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_941, x_961, x_942); +x_962 = l___private_Lake_CLI_Translate_Lean_0__Lake_addDeclFieldCore(x_941, x_961, x_943); x_928 = x_962; goto block_940; } @@ -30054,8 +30054,8 @@ if (x_972 == 0) { lean_dec(x_970); lean_dec_ref(x_969); -x_942 = x_964; -x_943 = x_968; +x_942 = x_968; +x_943 = x_964; x_944 = x_972; goto block_963; } @@ -30066,8 +30066,8 @@ x_973 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Lean_0__Lake_Lean lean_dec_ref(x_969); if (x_973 == 0) { -x_942 = x_964; -x_943 = x_968; +x_942 = x_968; +x_943 = x_964; x_944 = x_973; goto block_963; } @@ -30091,11 +30091,11 @@ goto block_974; { if (x_983 == 0) { -if (x_982 == 0) +if (x_981 == 0) { lean_object* x_984; x_984 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__1; -x_976 = x_981; +x_976 = x_982; x_977 = x_984; goto block_980; } @@ -30103,14 +30103,14 @@ else { lean_object* x_985; x_985 = l___private_Lake_CLI_Translate_Lean_0__Lake_instToLeanBoolMkStr1___lam__0___closed__4; -x_976 = x_981; +x_976 = x_982; x_977 = x_985; goto block_980; } } else { -x_964 = x_981; +x_964 = x_982; goto block_974; } } @@ -30142,8 +30142,8 @@ else uint8_t x_995; uint8_t x_996; x_995 = lean_unbox(x_991); x_996 = lean_unbox(x_991); -x_981 = x_987; -x_982 = x_995; +x_981 = x_995; +x_982 = x_987; x_983 = x_996; goto block_986; } @@ -30153,8 +30153,8 @@ else uint8_t x_997; uint8_t x_998; x_997 = lean_unbox(x_991); x_998 = lean_unbox(x_992); -x_981 = x_987; -x_982 = x_997; +x_981 = x_997; +x_982 = x_987; x_983 = x_998; goto block_986; } @@ -30588,7 +30588,7 @@ lean_ctor_set(x_20, 1, x_11); lean_ctor_set(x_20, 2, x_19); lean_inc(x_9); x_21 = l_Lean_Syntax_node1(x_9, x_15, x_20); -x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_17, x_16, x_14, x_21); +x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_17, x_14, x_16, x_21); if (lean_is_scalar(x_7)) { x_23 = lean_alloc_ctor(0, 2, 0); } else { @@ -30624,9 +30624,9 @@ if (lean_obj_tag(x_2) == 0) { lean_object* x_35; x_35 = l___private_Lake_CLI_Translate_Lean_0__Lake_Dependency_mkRequire___lam__0___closed__0; -x_14 = x_33; +x_14 = x_29; x_15 = x_34; -x_16 = x_29; +x_16 = x_33; x_17 = x_27; x_18 = x_35; goto block_24; @@ -30638,9 +30638,9 @@ x_36 = lean_ctor_get(x_2, 0); lean_inc(x_36); lean_dec_ref(x_2); x_37 = l_Array_mkArray1___redArg(x_36); -x_14 = x_33; +x_14 = x_29; x_15 = x_34; -x_16 = x_29; +x_16 = x_33; x_17 = x_27; x_18 = x_37; goto block_24; @@ -30980,8 +30980,8 @@ lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_11); lean_ctor_set(x_20, 2, x_19); lean_inc(x_9); -x_21 = l_Lean_Syntax_node1(x_9, x_15, x_20); -x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_14, x_16, x_17, x_21); +x_21 = l_Lean_Syntax_node1(x_9, x_17, x_20); +x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_16, x_15, x_14, x_21); if (lean_is_scalar(x_7)) { x_23 = lean_alloc_ctor(0, 2, 0); } else { @@ -31017,10 +31017,10 @@ if (lean_obj_tag(x_2) == 0) { lean_object* x_35; x_35 = l___private_Lake_CLI_Translate_Lean_0__Lake_Dependency_mkRequire___lam__0___closed__0; -x_14 = x_27; -x_15 = x_34; -x_16 = x_29; -x_17 = x_33; +x_14 = x_33; +x_15 = x_29; +x_16 = x_27; +x_17 = x_34; x_18 = x_35; goto block_24; } @@ -31031,10 +31031,10 @@ x_36 = lean_ctor_get(x_2, 0); lean_inc(x_36); lean_dec_ref(x_2); x_37 = l_Array_mkArray1___redArg(x_36); -x_14 = x_27; -x_15 = x_34; -x_16 = x_29; -x_17 = x_33; +x_14 = x_33; +x_15 = x_29; +x_16 = x_27; +x_17 = x_34; x_18 = x_37; goto block_24; } @@ -31164,7 +31164,7 @@ lean_ctor_set(x_20, 1, x_11); lean_ctor_set(x_20, 2, x_19); lean_inc(x_9); x_21 = l_Lean_Syntax_node1(x_9, x_16, x_20); -x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_15, x_14, x_17, x_21); +x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_14, x_17, x_15, x_21); if (lean_is_scalar(x_7)) { x_23 = lean_alloc_ctor(0, 2, 0); } else { @@ -31200,10 +31200,10 @@ if (lean_obj_tag(x_2) == 0) { lean_object* x_35; x_35 = l___private_Lake_CLI_Translate_Lean_0__Lake_Dependency_mkRequire___lam__0___closed__0; -x_14 = x_29; -x_15 = x_27; +x_14 = x_27; +x_15 = x_33; x_16 = x_34; -x_17 = x_33; +x_17 = x_29; x_18 = x_35; goto block_24; } @@ -31214,10 +31214,10 @@ x_36 = lean_ctor_get(x_2, 0); lean_inc(x_36); lean_dec_ref(x_2); x_37 = l_Array_mkArray1___redArg(x_36); -x_14 = x_29; -x_15 = x_27; +x_14 = x_27; +x_15 = x_33; x_16 = x_34; -x_17 = x_33; +x_17 = x_29; x_18 = x_37; goto block_24; } @@ -31346,8 +31346,8 @@ lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_11); lean_ctor_set(x_20, 2, x_19); lean_inc(x_9); -x_21 = l_Lean_Syntax_node1(x_9, x_17, x_20); -x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_16, x_14, x_15, x_21); +x_21 = l_Lean_Syntax_node1(x_9, x_16, x_20); +x_22 = l_Lean_Syntax_node5(x_9, x_10, x_13, x_15, x_14, x_17, x_21); if (lean_is_scalar(x_7)) { x_23 = lean_alloc_ctor(0, 2, 0); } else { @@ -31384,9 +31384,9 @@ if (lean_obj_tag(x_2) == 0) lean_object* x_35; x_35 = l___private_Lake_CLI_Translate_Lean_0__Lake_Dependency_mkRequire___lam__0___closed__0; x_14 = x_29; -x_15 = x_33; -x_16 = x_27; -x_17 = x_34; +x_15 = x_27; +x_16 = x_34; +x_17 = x_33; x_18 = x_35; goto block_24; } @@ -31398,9 +31398,9 @@ lean_inc(x_36); lean_dec_ref(x_2); x_37 = l_Array_mkArray1___redArg(x_36); x_14 = x_29; -x_15 = x_33; -x_16 = x_27; -x_17 = x_34; +x_15 = x_27; +x_16 = x_34; +x_17 = x_33; x_18 = x_37; goto block_24; } diff --git a/stage0/stdlib/Lake/CLI/Translate/Toml.c b/stage0/stdlib/Lake/CLI/Translate/Toml.c index ca819c87689c..bcfe372d434e 100644 --- a/stage0/stdlib/Lake/CLI/Translate/Toml.c +++ b/stage0/stdlib/Lake/CLI/Translate/Toml.c @@ -4815,8 +4815,8 @@ x_20 = l_Lean_Name_str___override(x_2, x_19); x_21 = l_Lean_Name_append(x_18, x_20); x_22 = 0; x_23 = l_Lean_mkIdentFromRef___redArg(x_7, x_8, x_21, x_22); -lean_inc_ref(x_15); -x_24 = lean_apply_2(x_23, x_15, x_17); +lean_inc_ref(x_13); +x_24 = lean_apply_2(x_23, x_13, x_17); 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; @@ -4830,8 +4830,8 @@ lean_inc(x_2); x_28 = l_Lean_Name_str___override(x_2, x_27); x_29 = l_Lean_Name_append(x_18, x_28); x_30 = l_Lean_mkIdentFromRef___redArg(x_7, x_8, x_29, x_22); -lean_inc_ref(x_15); -x_31 = lean_apply_2(x_30, x_15, x_26); +lean_inc_ref(x_13); +x_31 = lean_apply_2(x_30, x_13, x_26); if (lean_obj_tag(x_31) == 0) { uint8_t x_32; @@ -4840,17 +4840,17 @@ if (x_32 == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; x_33 = lean_ctor_get(x_31, 0); -x_34 = lean_ctor_get(x_15, 1); +x_34 = lean_ctor_get(x_13, 1); lean_inc(x_34); -x_35 = lean_ctor_get(x_15, 2); +x_35 = lean_ctor_get(x_13, 2); lean_inc(x_35); -x_36 = lean_ctor_get(x_15, 5); +x_36 = lean_ctor_get(x_13, 5); lean_inc(x_36); -lean_dec_ref(x_15); -x_37 = lean_mk_empty_array_with_capacity(x_13); -lean_inc(x_13); -x_38 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop(lean_box(0), x_13, x_12, x_13, lean_box(0), x_37); -lean_dec(x_13); +lean_dec_ref(x_13); +x_37 = lean_mk_empty_array_with_capacity(x_15); +lean_inc(x_15); +x_38 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop(lean_box(0), x_15, x_12, x_15, lean_box(0), x_37); +lean_dec(x_15); x_39 = l_Lean_Syntax_mkCApp(x_2, x_38); x_40 = l_Lean_SourceInfo_fromRef(x_36, x_22); lean_dec(x_36); @@ -5151,17 +5151,17 @@ x_171 = lean_ctor_get(x_31, 1); lean_inc(x_171); lean_inc(x_170); lean_dec(x_31); -x_172 = lean_ctor_get(x_15, 1); +x_172 = lean_ctor_get(x_13, 1); lean_inc(x_172); -x_173 = lean_ctor_get(x_15, 2); +x_173 = lean_ctor_get(x_13, 2); lean_inc(x_173); -x_174 = lean_ctor_get(x_15, 5); +x_174 = lean_ctor_get(x_13, 5); lean_inc(x_174); -lean_dec_ref(x_15); -x_175 = lean_mk_empty_array_with_capacity(x_13); -lean_inc(x_13); -x_176 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop(lean_box(0), x_13, x_12, x_13, lean_box(0), x_175); -lean_dec(x_13); +lean_dec_ref(x_13); +x_175 = lean_mk_empty_array_with_capacity(x_15); +lean_inc(x_15); +x_176 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop(lean_box(0), x_15, x_12, x_15, lean_box(0), x_175); +lean_dec(x_15); x_177 = l_Lean_Syntax_mkCApp(x_2, x_176); x_178 = l_Lean_SourceInfo_fromRef(x_174, x_22); lean_dec(x_174); @@ -5462,8 +5462,8 @@ else uint8_t x_309; lean_dec(x_25); lean_dec(x_16); -lean_dec_ref(x_15); -lean_dec(x_13); +lean_dec(x_15); +lean_dec_ref(x_13); lean_dec_ref(x_12); lean_dec(x_11); lean_dec(x_2); @@ -5492,8 +5492,8 @@ else { uint8_t x_313; lean_dec(x_16); -lean_dec_ref(x_15); -lean_dec(x_13); +lean_dec(x_15); +lean_dec_ref(x_13); lean_dec_ref(x_12); lean_dec(x_11); lean_dec(x_2); @@ -5529,9 +5529,9 @@ if (x_325 == 0) lean_dec(x_324); lean_dec_ref(x_319); lean_dec_ref(x_9); -x_13 = x_10; +x_13 = x_321; x_14 = x_323; -x_15 = x_321; +x_15 = x_10; x_16 = x_320; x_17 = x_322; goto block_317; @@ -5545,9 +5545,9 @@ if (x_326 == 0) lean_dec(x_324); lean_dec_ref(x_319); lean_dec_ref(x_9); -x_13 = x_10; +x_13 = x_321; x_14 = x_323; -x_15 = x_321; +x_15 = x_10; x_16 = x_320; x_17 = x_322; goto block_317; @@ -5569,9 +5569,9 @@ lean_inc(x_331); x_332 = lean_ctor_get(x_330, 1); lean_inc(x_332); lean_dec_ref(x_330); -x_13 = x_10; +x_13 = x_321; x_14 = x_323; -x_15 = x_321; +x_15 = x_10; x_16 = x_331; x_17 = x_332; goto block_317; @@ -6076,7 +6076,7 @@ lean_inc(x_2); x_13 = l_Lean_Name_str___override(x_2, x_12); x_14 = l_Lean_Name_append(x_11, x_13); x_15 = 0; -x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_6, x_10); +x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_8, x_10); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -6087,23 +6087,23 @@ x_20 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_21 = l_Lean_Name_str___override(x_2, x_20); x_22 = l_Lean_Name_append(x_11, x_21); -x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_6, x_19); +x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_8, x_19); 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_6, 1); +x_26 = lean_ctor_get(x_8, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_6, 2); +x_27 = lean_ctor_get(x_8, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_6, 5); +x_28 = lean_ctor_get(x_8, 5); lean_inc(x_28); -lean_dec_ref(x_6); -x_29 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); -lean_dec(x_7); +lean_dec_ref(x_8); +x_29 = lean_mk_empty_array_with_capacity(x_6); +lean_inc(x_6); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_29); +lean_dec(x_6); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -6126,7 +6126,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_8); +x_42 = lean_mk_empty_array_with_capacity(x_7); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -6399,17 +6399,17 @@ x_162 = lean_ctor_get(x_23, 1); lean_inc(x_162); lean_inc(x_161); lean_dec(x_23); -x_163 = lean_ctor_get(x_6, 1); +x_163 = lean_ctor_get(x_8, 1); lean_inc(x_163); -x_164 = lean_ctor_get(x_6, 2); +x_164 = lean_ctor_get(x_8, 2); lean_inc(x_164); -x_165 = lean_ctor_get(x_6, 5); +x_165 = lean_ctor_get(x_8, 5); lean_inc(x_165); -lean_dec_ref(x_6); -x_166 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); -lean_dec(x_7); +lean_dec_ref(x_8); +x_166 = lean_mk_empty_array_with_capacity(x_6); +lean_inc(x_6); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_166); +lean_dec(x_6); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -6432,7 +6432,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_8); +x_179 = lean_mk_empty_array_with_capacity(x_7); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -6712,7 +6712,7 @@ x_301 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_302 = l_Lean_Name_str___override(x_2, x_301); x_303 = l_Lean_Name_append(x_11, x_302); -x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_6, x_300); +x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_8, x_300); x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); x_306 = lean_ctor_get(x_304, 1); @@ -6725,17 +6725,17 @@ if (lean_is_exclusive(x_304)) { lean_dec_ref(x_304); x_307 = lean_box(0); } -x_308 = lean_ctor_get(x_6, 1); +x_308 = lean_ctor_get(x_8, 1); lean_inc(x_308); -x_309 = lean_ctor_get(x_6, 2); +x_309 = lean_ctor_get(x_8, 2); lean_inc(x_309); -x_310 = lean_ctor_get(x_6, 5); +x_310 = lean_ctor_get(x_8, 5); lean_inc(x_310); -lean_dec_ref(x_6); -x_311 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); -lean_dec(x_7); +lean_dec_ref(x_8); +x_311 = lean_mk_empty_array_with_capacity(x_6); +lean_inc(x_6); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_311); +lean_dec(x_6); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -6758,7 +6758,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_8); +x_325 = lean_mk_empty_array_with_capacity(x_7); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -7044,9 +7044,9 @@ if (x_454 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_6 = x_451; +x_7 = x_452; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -7059,9 +7059,9 @@ if (x_455 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_6 = x_451; +x_7 = x_452; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -7080,9 +7080,9 @@ lean_inc(x_459); x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); -x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_6 = x_451; +x_7 = x_452; +x_8 = x_448; x_9 = x_459; x_10 = x_460; goto block_445; @@ -7181,7 +7181,7 @@ lean_inc(x_2); x_13 = l_Lean_Name_str___override(x_2, x_12); x_14 = l_Lean_Name_append(x_11, x_13); x_15 = 0; -x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_7, x_10); +x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_8, x_10); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -7192,23 +7192,23 @@ x_20 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_21 = l_Lean_Name_str___override(x_2, x_20); x_22 = l_Lean_Name_append(x_11, x_21); -x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_7, x_19); +x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_8, x_19); 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_7, 1); +x_26 = lean_ctor_get(x_8, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_7, 2); +x_27 = lean_ctor_get(x_8, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_7, 5); +x_28 = lean_ctor_get(x_8, 5); lean_inc(x_28); -lean_dec_ref(x_7); -x_29 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_29); -lean_dec(x_6); +lean_dec_ref(x_8); +x_29 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); +lean_dec(x_7); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -7231,7 +7231,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_8); +x_42 = lean_mk_empty_array_with_capacity(x_6); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -7504,17 +7504,17 @@ x_162 = lean_ctor_get(x_23, 1); lean_inc(x_162); lean_inc(x_161); lean_dec(x_23); -x_163 = lean_ctor_get(x_7, 1); +x_163 = lean_ctor_get(x_8, 1); lean_inc(x_163); -x_164 = lean_ctor_get(x_7, 2); +x_164 = lean_ctor_get(x_8, 2); lean_inc(x_164); -x_165 = lean_ctor_get(x_7, 5); +x_165 = lean_ctor_get(x_8, 5); lean_inc(x_165); -lean_dec_ref(x_7); -x_166 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_166); -lean_dec(x_6); +lean_dec_ref(x_8); +x_166 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); +lean_dec(x_7); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -7537,7 +7537,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_8); +x_179 = lean_mk_empty_array_with_capacity(x_6); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -7817,7 +7817,7 @@ x_301 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_302 = l_Lean_Name_str___override(x_2, x_301); x_303 = l_Lean_Name_append(x_11, x_302); -x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_7, x_300); +x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_8, x_300); x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); x_306 = lean_ctor_get(x_304, 1); @@ -7830,17 +7830,17 @@ if (lean_is_exclusive(x_304)) { lean_dec_ref(x_304); x_307 = lean_box(0); } -x_308 = lean_ctor_get(x_7, 1); +x_308 = lean_ctor_get(x_8, 1); lean_inc(x_308); -x_309 = lean_ctor_get(x_7, 2); +x_309 = lean_ctor_get(x_8, 2); lean_inc(x_309); -x_310 = lean_ctor_get(x_7, 5); +x_310 = lean_ctor_get(x_8, 5); lean_inc(x_310); -lean_dec_ref(x_7); -x_311 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_311); -lean_dec(x_6); +lean_dec_ref(x_8); +x_311 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); +lean_dec(x_7); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -7863,7 +7863,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_8); +x_325 = lean_mk_empty_array_with_capacity(x_6); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -8149,9 +8149,9 @@ if (x_454 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -8164,9 +8164,9 @@ if (x_455 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -8185,9 +8185,9 @@ lean_inc(x_459); x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_459; x_10 = x_460; goto block_445; @@ -8286,7 +8286,7 @@ lean_inc(x_2); x_13 = l_Lean_Name_str___override(x_2, x_12); x_14 = l_Lean_Name_append(x_11, x_13); x_15 = 0; -x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_7, x_10); +x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_8, x_10); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -8297,23 +8297,23 @@ x_20 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_21 = l_Lean_Name_str___override(x_2, x_20); x_22 = l_Lean_Name_append(x_11, x_21); -x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_7, x_19); +x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_8, x_19); 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_7, 1); +x_26 = lean_ctor_get(x_8, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_7, 2); +x_27 = lean_ctor_get(x_8, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_7, 5); +x_28 = lean_ctor_get(x_8, 5); lean_inc(x_28); -lean_dec_ref(x_7); -x_29 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_29); -lean_dec(x_6); +lean_dec_ref(x_8); +x_29 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); +lean_dec(x_7); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -8336,7 +8336,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_8); +x_42 = lean_mk_empty_array_with_capacity(x_6); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -8609,17 +8609,17 @@ x_162 = lean_ctor_get(x_23, 1); lean_inc(x_162); lean_inc(x_161); lean_dec(x_23); -x_163 = lean_ctor_get(x_7, 1); +x_163 = lean_ctor_get(x_8, 1); lean_inc(x_163); -x_164 = lean_ctor_get(x_7, 2); +x_164 = lean_ctor_get(x_8, 2); lean_inc(x_164); -x_165 = lean_ctor_get(x_7, 5); +x_165 = lean_ctor_get(x_8, 5); lean_inc(x_165); -lean_dec_ref(x_7); -x_166 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_166); -lean_dec(x_6); +lean_dec_ref(x_8); +x_166 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); +lean_dec(x_7); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -8642,7 +8642,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_8); +x_179 = lean_mk_empty_array_with_capacity(x_6); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -8922,7 +8922,7 @@ x_301 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_302 = l_Lean_Name_str___override(x_2, x_301); x_303 = l_Lean_Name_append(x_11, x_302); -x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_7, x_300); +x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_8, x_300); x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); x_306 = lean_ctor_get(x_304, 1); @@ -8935,17 +8935,17 @@ if (lean_is_exclusive(x_304)) { lean_dec_ref(x_304); x_307 = lean_box(0); } -x_308 = lean_ctor_get(x_7, 1); +x_308 = lean_ctor_get(x_8, 1); lean_inc(x_308); -x_309 = lean_ctor_get(x_7, 2); +x_309 = lean_ctor_get(x_8, 2); lean_inc(x_309); -x_310 = lean_ctor_get(x_7, 5); +x_310 = lean_ctor_get(x_8, 5); lean_inc(x_310); -lean_dec_ref(x_7); -x_311 = lean_mk_empty_array_with_capacity(x_6); -lean_inc(x_6); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_6, x_6, x_311); -lean_dec(x_6); +lean_dec_ref(x_8); +x_311 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); +lean_dec(x_7); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -8968,7 +8968,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_8); +x_325 = lean_mk_empty_array_with_capacity(x_6); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -9254,9 +9254,9 @@ if (x_454 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -9269,9 +9269,9 @@ if (x_455 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_447; x_10 = x_449; goto block_445; @@ -9290,9 +9290,9 @@ lean_inc(x_459); x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); -x_6 = x_451; -x_7 = x_448; -x_8 = x_452; +x_6 = x_452; +x_7 = x_451; +x_8 = x_448; x_9 = x_459; x_10 = x_460; goto block_445; @@ -9415,10 +9415,10 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_6, 5); lean_inc(x_28); lean_dec_ref(x_6); -x_29 = lean_mk_empty_array_with_capacity(x_8); -lean_inc(x_8); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_29); -lean_dec(x_8); +x_29 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); +lean_dec(x_7); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -9441,7 +9441,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_7); +x_42 = lean_mk_empty_array_with_capacity(x_8); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -9721,10 +9721,10 @@ lean_inc(x_164); x_165 = lean_ctor_get(x_6, 5); lean_inc(x_165); lean_dec_ref(x_6); -x_166 = lean_mk_empty_array_with_capacity(x_8); -lean_inc(x_8); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_166); -lean_dec(x_8); +x_166 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); +lean_dec(x_7); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -9747,7 +9747,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_7); +x_179 = lean_mk_empty_array_with_capacity(x_8); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -10047,10 +10047,10 @@ lean_inc(x_309); x_310 = lean_ctor_get(x_6, 5); lean_inc(x_310); lean_dec_ref(x_6); -x_311 = lean_mk_empty_array_with_capacity(x_8); -lean_inc(x_8); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_311); -lean_dec(x_8); +x_311 = lean_mk_empty_array_with_capacity(x_7); +lean_inc(x_7); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); +lean_dec(x_7); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -10073,7 +10073,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_7); +x_325 = lean_mk_empty_array_with_capacity(x_8); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -10360,8 +10360,8 @@ if (x_454 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_452; -x_8 = x_451; +x_7 = x_451; +x_8 = x_452; x_9 = x_447; x_10 = x_449; goto block_445; @@ -10375,8 +10375,8 @@ if (x_455 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_452; -x_8 = x_451; +x_7 = x_451; +x_8 = x_452; x_9 = x_447; x_10 = x_449; goto block_445; @@ -10396,8 +10396,8 @@ x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); x_6 = x_448; -x_7 = x_452; -x_8 = x_451; +x_7 = x_451; +x_8 = x_452; x_9 = x_459; x_10 = x_460; goto block_445; @@ -10520,10 +10520,10 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_6, 5); lean_inc(x_28); lean_dec_ref(x_6); -x_29 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); -lean_dec(x_7); +x_29 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_29); +lean_dec(x_8); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -10546,7 +10546,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_8); +x_42 = lean_mk_empty_array_with_capacity(x_7); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -10826,10 +10826,10 @@ lean_inc(x_164); x_165 = lean_ctor_get(x_6, 5); lean_inc(x_165); lean_dec_ref(x_6); -x_166 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); -lean_dec(x_7); +x_166 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_166); +lean_dec(x_8); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -10852,7 +10852,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_8); +x_179 = lean_mk_empty_array_with_capacity(x_7); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -11152,10 +11152,10 @@ lean_inc(x_309); x_310 = lean_ctor_get(x_6, 5); lean_inc(x_310); lean_dec_ref(x_6); -x_311 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); -lean_dec(x_7); +x_311 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_311); +lean_dec(x_8); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -11178,7 +11178,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_8); +x_325 = lean_mk_empty_array_with_capacity(x_7); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -11465,8 +11465,8 @@ if (x_454 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -11480,8 +11480,8 @@ if (x_455 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -11501,8 +11501,8 @@ x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_459; x_10 = x_460; goto block_445; @@ -11601,7 +11601,7 @@ lean_inc(x_2); x_13 = l_Lean_Name_str___override(x_2, x_12); x_14 = l_Lean_Name_append(x_11, x_13); x_15 = 0; -x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_8, x_10); +x_16 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_14, x_15, x_6, x_10); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -11612,23 +11612,23 @@ x_20 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_21 = l_Lean_Name_str___override(x_2, x_20); x_22 = l_Lean_Name_append(x_11, x_21); -x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_8, x_19); +x_23 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_22, x_15, x_6, x_19); 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_8, 1); +x_26 = lean_ctor_get(x_6, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_8, 2); +x_27 = lean_ctor_get(x_6, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_8, 5); +x_28 = lean_ctor_get(x_6, 5); lean_inc(x_28); -lean_dec_ref(x_8); -x_29 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); -lean_dec(x_7); +lean_dec_ref(x_6); +x_29 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_29); +lean_dec(x_8); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -11651,7 +11651,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_6); +x_42 = lean_mk_empty_array_with_capacity(x_7); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -11924,17 +11924,17 @@ x_162 = lean_ctor_get(x_23, 1); lean_inc(x_162); lean_inc(x_161); lean_dec(x_23); -x_163 = lean_ctor_get(x_8, 1); +x_163 = lean_ctor_get(x_6, 1); lean_inc(x_163); -x_164 = lean_ctor_get(x_8, 2); +x_164 = lean_ctor_get(x_6, 2); lean_inc(x_164); -x_165 = lean_ctor_get(x_8, 5); +x_165 = lean_ctor_get(x_6, 5); lean_inc(x_165); -lean_dec_ref(x_8); -x_166 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); -lean_dec(x_7); +lean_dec_ref(x_6); +x_166 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_166); +lean_dec(x_8); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -11957,7 +11957,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_6); +x_179 = lean_mk_empty_array_with_capacity(x_7); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -12237,7 +12237,7 @@ x_301 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__14; lean_inc(x_2); x_302 = l_Lean_Name_str___override(x_2, x_301); x_303 = l_Lean_Name_append(x_11, x_302); -x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_8, x_300); +x_304 = l_Lean_mkIdentFromRef___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__2(x_303, x_15, x_6, x_300); x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); x_306 = lean_ctor_get(x_304, 1); @@ -12250,17 +12250,17 @@ if (lean_is_exclusive(x_304)) { lean_dec_ref(x_304); x_307 = lean_box(0); } -x_308 = lean_ctor_get(x_8, 1); +x_308 = lean_ctor_get(x_6, 1); lean_inc(x_308); -x_309 = lean_ctor_get(x_8, 2); +x_309 = lean_ctor_get(x_6, 2); lean_inc(x_309); -x_310 = lean_ctor_get(x_8, 5); +x_310 = lean_ctor_get(x_6, 5); lean_inc(x_310); -lean_dec_ref(x_8); -x_311 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); -lean_dec(x_7); +lean_dec_ref(x_6); +x_311 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_311); +lean_dec(x_8); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -12283,7 +12283,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_6); +x_325 = lean_mk_empty_array_with_capacity(x_7); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -12569,9 +12569,9 @@ if (x_454 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_452; -x_7 = x_451; -x_8 = x_448; +x_6 = x_448; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -12584,9 +12584,9 @@ if (x_455 == 0) { lean_dec(x_453); lean_dec_ref(x_450); -x_6 = x_452; -x_7 = x_451; -x_8 = x_448; +x_6 = x_448; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -12605,9 +12605,9 @@ lean_inc(x_459); x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); -x_6 = x_452; -x_7 = x_451; -x_8 = x_448; +x_6 = x_448; +x_7 = x_452; +x_8 = x_451; x_9 = x_459; x_10 = x_460; goto block_445; @@ -12730,10 +12730,10 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_6, 5); lean_inc(x_28); lean_dec_ref(x_6); -x_29 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_29); -lean_dec(x_7); +x_29 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_30 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_29); +lean_dec(x_8); x_31 = l_Lean_Syntax_mkCApp(x_2, x_30); x_32 = l_Lean_SourceInfo_fromRef(x_28, x_15); lean_dec(x_28); @@ -12756,7 +12756,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_40); lean_ctor_set(x_16, 0, x_32); x_41 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_42 = lean_mk_empty_array_with_capacity(x_8); +x_42 = lean_mk_empty_array_with_capacity(x_7); x_43 = lean_box(2); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_43); @@ -13036,10 +13036,10 @@ lean_inc(x_164); x_165 = lean_ctor_get(x_6, 5); lean_inc(x_165); lean_dec_ref(x_6); -x_166 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_166); -lean_dec(x_7); +x_166 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_167 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_166); +lean_dec(x_8); x_168 = l_Lean_Syntax_mkCApp(x_2, x_167); x_169 = l_Lean_SourceInfo_fromRef(x_165, x_15); lean_dec(x_165); @@ -13062,7 +13062,7 @@ lean_ctor_set_tag(x_16, 2); lean_ctor_set(x_16, 1, x_177); lean_ctor_set(x_16, 0, x_169); x_178 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_179 = lean_mk_empty_array_with_capacity(x_8); +x_179 = lean_mk_empty_array_with_capacity(x_7); x_180 = lean_box(2); x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_180); @@ -13362,10 +13362,10 @@ lean_inc(x_309); x_310 = lean_ctor_get(x_6, 5); lean_inc(x_310); lean_dec_ref(x_6); -x_311 = lean_mk_empty_array_with_capacity(x_7); -lean_inc(x_7); -x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_7, x_7, x_311); -lean_dec(x_7); +x_311 = lean_mk_empty_array_with_capacity(x_8); +lean_inc(x_8); +x_312 = l___private_Init_Data_Nat_Fold_0__Nat_foldTR_loop___at___00__private_Lake_CLI_Translate_Toml_0__Lake_genToToml___at___00__private_Lake_CLI_Translate_Toml_0__Lake___aux__Lake__CLI__Translate__Toml______macroRules____private__Lake__CLI__Translate__Toml__0__Lake__commandGen__toml__encoders_x25__1_spec__0_spec__3___redArg(x_8, x_8, x_311); +lean_dec(x_8); x_313 = l_Lean_Syntax_mkCApp(x_2, x_312); x_314 = l_Lean_SourceInfo_fromRef(x_310, x_15); lean_dec(x_310); @@ -13388,7 +13388,7 @@ x_323 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_323, 0, x_314); lean_ctor_set(x_323, 1, x_322); x_324 = l___private_Lake_CLI_Translate_Toml_0__Lake_genToToml___closed__25; -x_325 = lean_mk_empty_array_with_capacity(x_8); +x_325 = lean_mk_empty_array_with_capacity(x_7); x_326 = lean_box(2); x_327 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_327, 0, x_326); @@ -13675,8 +13675,8 @@ if (x_454 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -13690,8 +13690,8 @@ if (x_455 == 0) lean_dec(x_453); lean_dec_ref(x_450); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_447; x_10 = x_449; goto block_445; @@ -13711,8 +13711,8 @@ x_460 = lean_ctor_get(x_458, 1); lean_inc(x_460); lean_dec_ref(x_458); x_6 = x_448; -x_7 = x_451; -x_8 = x_452; +x_7 = x_452; +x_8 = x_451; x_9 = x_459; x_10 = x_460; goto block_445; @@ -14784,13 +14784,13 @@ goto block_45; lean_object* x_66; lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; x_66 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_67 = lean_box(0); -x_68 = lean_array_size(x_65); +x_68 = lean_array_size(x_64); x_69 = 0; -x_70 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_68, x_69, x_65); +x_70 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_68, x_69, x_64); x_71 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_71, 0, x_67); lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lake_Toml_RBDict_insert___redArg(x_66, x_63, x_71, x_64); +x_72 = l_Lake_Toml_RBDict_insert___redArg(x_66, x_63, x_71, x_65); x_47 = x_72; goto block_62; } @@ -14814,8 +14814,8 @@ if (x_82 == 0) { lean_dec(x_80); lean_dec_ref(x_79); -x_64 = x_74; -x_65 = x_78; +x_64 = x_78; +x_65 = x_74; goto block_73; } else @@ -14825,8 +14825,8 @@ x_83 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanC lean_dec_ref(x_79); if (x_83 == 0) { -x_64 = x_74; -x_65 = x_78; +x_64 = x_78; +x_65 = x_74; goto block_73; } else @@ -14842,13 +14842,13 @@ goto block_62; lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; x_88 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_89 = lean_box(0); -x_90 = lean_array_size(x_87); +x_90 = lean_array_size(x_86); x_91 = 0; -x_92 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_90, x_91, x_87); +x_92 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_90, x_91, x_86); x_93 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_93, 0, x_89); lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lake_Toml_RBDict_insert___redArg(x_88, x_85, x_93, x_86); +x_94 = l_Lake_Toml_RBDict_insert___redArg(x_88, x_85, x_93, x_87); x_74 = x_94; goto block_84; } @@ -14872,8 +14872,8 @@ if (x_104 == 0) { lean_dec(x_102); lean_dec_ref(x_101); -x_86 = x_96; -x_87 = x_100; +x_86 = x_100; +x_87 = x_96; goto block_95; } else @@ -14883,8 +14883,8 @@ x_105 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_101); if (x_105 == 0) { -x_86 = x_96; -x_87 = x_100; +x_86 = x_100; +x_87 = x_96; goto block_95; } else @@ -15943,13 +15943,13 @@ goto block_75; lean_object* x_104; lean_object* x_105; size_t x_106; size_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; x_104 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_105 = lean_box(0); -x_106 = lean_array_size(x_102); +x_106 = lean_array_size(x_103); x_107 = 0; -x_108 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_106, x_107, x_102); +x_108 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_106, x_107, x_103); x_109 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_109, 0, x_105); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lake_Toml_RBDict_insert___redArg(x_104, x_101, x_109, x_103); +x_110 = l_Lake_Toml_RBDict_insert___redArg(x_104, x_101, x_109, x_102); x_87 = x_110; goto block_100; } @@ -15978,8 +15978,8 @@ if (x_123 == 0) { lean_dec(x_121); lean_dec_ref(x_120); -x_102 = x_119; -x_103 = x_112; +x_102 = x_112; +x_103 = x_119; goto block_111; } else @@ -15989,8 +15989,8 @@ x_124 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_120); if (x_124 == 0) { -x_102 = x_119; -x_103 = x_112; +x_102 = x_112; +x_103 = x_119; goto block_111; } else @@ -16239,13 +16239,13 @@ goto block_199; 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_object* x_234; x_228 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_229 = lean_box(0); -x_230 = lean_array_size(x_226); +x_230 = lean_array_size(x_227); x_231 = 0; -x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_230, x_231, x_226); +x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_230, x_231, x_227); x_233 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_233, 0, x_229); lean_ctor_set(x_233, 1, x_232); -x_234 = l_Lake_Toml_RBDict_insert___redArg(x_228, x_225, x_233, x_227); +x_234 = l_Lake_Toml_RBDict_insert___redArg(x_228, x_225, x_233, x_226); x_211 = x_234; goto block_224; } @@ -16274,8 +16274,8 @@ if (x_247 == 0) { lean_dec(x_245); lean_dec_ref(x_244); -x_226 = x_243; -x_227 = x_236; +x_226 = x_236; +x_227 = x_243; goto block_235; } else @@ -16285,8 +16285,8 @@ x_248 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_244); if (x_248 == 0) { -x_226 = x_243; -x_227 = x_236; +x_226 = x_236; +x_227 = x_243; goto block_235; } else @@ -16623,13 +16623,13 @@ goto block_358; lean_object* x_375; lean_object* x_376; size_t x_377; size_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; x_375 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_376 = lean_box(0); -x_377 = lean_array_size(x_373); +x_377 = lean_array_size(x_374); x_378 = 0; -x_379 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__1(x_377, x_378, x_373); +x_379 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__1(x_377, x_378, x_374); x_380 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_380, 0, x_376); lean_ctor_set(x_380, 1, x_379); -x_381 = l_Lake_Toml_RBDict_insert___redArg(x_375, x_372, x_380, x_374); +x_381 = l_Lake_Toml_RBDict_insert___redArg(x_375, x_372, x_380, x_373); x_359 = x_381; goto block_371; } @@ -16654,8 +16654,8 @@ if (x_391 == 0) { lean_dec(x_389); lean_dec_ref(x_388); -x_373 = x_387; -x_374 = x_383; +x_373 = x_383; +x_374 = x_387; goto block_382; } else @@ -16665,8 +16665,8 @@ x_392 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_388); if (x_392 == 0) { -x_373 = x_387; -x_374 = x_383; +x_373 = x_383; +x_374 = x_387; goto block_382; } else @@ -16814,13 +16814,13 @@ goto block_429; lean_object* x_446; lean_object* x_447; size_t x_448; size_t x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; x_446 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_447 = lean_box(0); -x_448 = lean_array_size(x_444); +x_448 = lean_array_size(x_445); x_449 = 0; -x_450 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__3(x_448, x_449, x_444); +x_450 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__3(x_448, x_449, x_445); x_451 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_451, 0, x_447); lean_ctor_set(x_451, 1, x_450); -x_452 = l_Lake_Toml_RBDict_insert___redArg(x_446, x_443, x_451, x_445); +x_452 = l_Lake_Toml_RBDict_insert___redArg(x_446, x_443, x_451, x_444); x_431 = x_452; goto block_442; } @@ -16845,8 +16845,8 @@ if (x_462 == 0) { lean_dec(x_460); lean_dec_ref(x_459); -x_444 = x_458; -x_445 = x_454; +x_444 = x_454; +x_445 = x_458; goto block_453; } else @@ -16856,8 +16856,8 @@ x_463 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_459); if (x_463 == 0) { -x_444 = x_458; -x_445 = x_454; +x_444 = x_454; +x_445 = x_458; goto block_453; } else @@ -17302,13 +17302,13 @@ goto block_55; lean_object* x_79; lean_object* x_80; size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; x_79 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_80 = lean_box(0); -x_81 = lean_array_size(x_77); +x_81 = lean_array_size(x_78); x_82 = 0; -x_83 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_81, x_82, x_77); +x_83 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_81, x_82, x_78); x_84 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_84, 0, x_80); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lake_Toml_RBDict_insert___redArg(x_79, x_76, x_84, x_78); +x_85 = l_Lake_Toml_RBDict_insert___redArg(x_79, x_76, x_84, x_77); x_57 = x_85; goto block_75; } @@ -17337,8 +17337,8 @@ if (x_98 == 0) { lean_dec(x_96); lean_dec_ref(x_95); -x_77 = x_94; -x_78 = x_87; +x_77 = x_87; +x_78 = x_94; goto block_86; } else @@ -17348,8 +17348,8 @@ x_99 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanC lean_dec_ref(x_95); if (x_99 == 0) { -x_77 = x_94; -x_78 = x_87; +x_77 = x_87; +x_78 = x_94; goto block_86; } else @@ -17500,13 +17500,13 @@ goto block_142; lean_object* x_163; lean_object* x_164; size_t x_165; size_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; x_163 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_164 = lean_box(0); -x_165 = lean_array_size(x_161); +x_165 = lean_array_size(x_162); x_166 = 0; -x_167 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_165, x_166, x_161); +x_167 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_165, x_166, x_162); x_168 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_168, 0, x_164); lean_ctor_set(x_168, 1, x_167); -x_169 = l_Lake_Toml_RBDict_insert___redArg(x_163, x_160, x_168, x_162); +x_169 = l_Lake_Toml_RBDict_insert___redArg(x_163, x_160, x_168, x_161); x_144 = x_169; goto block_159; } @@ -17535,8 +17535,8 @@ if (x_182 == 0) { lean_dec(x_180); lean_dec_ref(x_179); -x_161 = x_178; -x_162 = x_171; +x_161 = x_171; +x_162 = x_178; goto block_170; } else @@ -17546,8 +17546,8 @@ x_183 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_179); if (x_183 == 0) { -x_161 = x_178; -x_162 = x_171; +x_161 = x_171; +x_162 = x_178; goto block_170; } else @@ -17661,13 +17661,13 @@ goto block_199; 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_object* x_234; x_228 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_229 = lean_box(0); -x_230 = lean_array_size(x_227); +x_230 = lean_array_size(x_226); x_231 = 0; -x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_230, x_231, x_227); +x_232 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_230, x_231, x_226); x_233 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_233, 0, x_229); lean_ctor_set(x_233, 1, x_232); -x_234 = l_Lake_Toml_RBDict_insert___redArg(x_228, x_225, x_233, x_226); +x_234 = l_Lake_Toml_RBDict_insert___redArg(x_228, x_225, x_233, x_227); x_211 = x_234; goto block_224; } @@ -17696,8 +17696,8 @@ if (x_247 == 0) { lean_dec(x_245); lean_dec_ref(x_244); -x_226 = x_236; -x_227 = x_243; +x_226 = x_243; +x_227 = x_236; goto block_235; } else @@ -17707,8 +17707,8 @@ x_248 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_244); if (x_248 == 0) { -x_226 = x_236; -x_227 = x_243; +x_226 = x_243; +x_227 = x_236; goto block_235; } else @@ -17724,13 +17724,13 @@ goto block_224; lean_object* x_253; lean_object* x_254; size_t x_255; size_t x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; x_253 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_254 = lean_box(0); -x_255 = lean_array_size(x_252); +x_255 = lean_array_size(x_251); x_256 = 0; -x_257 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_255, x_256, x_252); +x_257 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_255, x_256, x_251); 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_Lake_Toml_RBDict_insert___redArg(x_253, x_250, x_258, x_251); +x_259 = l_Lake_Toml_RBDict_insert___redArg(x_253, x_250, x_258, x_252); x_236 = x_259; goto block_249; } @@ -17759,8 +17759,8 @@ if (x_272 == 0) { lean_dec(x_270); lean_dec_ref(x_269); -x_251 = x_261; -x_252 = x_268; +x_251 = x_268; +x_252 = x_261; goto block_260; } else @@ -17770,8 +17770,8 @@ x_273 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_269); if (x_273 == 0) { -x_251 = x_261; -x_252 = x_268; +x_251 = x_268; +x_252 = x_261; goto block_260; } else @@ -17927,13 +17927,13 @@ goto block_318; lean_object* x_335; lean_object* x_336; size_t x_337; size_t x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; x_335 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_336 = lean_box(0); -x_337 = lean_array_size(x_334); +x_337 = lean_array_size(x_333); x_338 = 0; -x_339 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__1(x_337, x_338, x_334); +x_339 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_LeanLibConfig_toToml_spec__1(x_337, x_338, x_333); x_340 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_340, 0, x_336); lean_ctor_set(x_340, 1, x_339); -x_341 = l_Lake_Toml_RBDict_insert___redArg(x_335, x_332, x_340, x_333); +x_341 = l_Lake_Toml_RBDict_insert___redArg(x_335, x_332, x_340, x_334); x_319 = x_341; goto block_331; } @@ -17958,8 +17958,8 @@ if (x_351 == 0) { lean_dec(x_349); lean_dec_ref(x_348); -x_333 = x_343; -x_334 = x_347; +x_333 = x_347; +x_334 = x_343; goto block_342; } else @@ -17969,8 +17969,8 @@ x_352 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_348); if (x_352 == 0) { -x_333 = x_343; -x_334 = x_347; +x_333 = x_347; +x_334 = x_343; goto block_342; } else @@ -18129,7 +18129,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lake_CLI_Translate_Toml_0__Lake_InputFileConfig_toToml(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; uint8_t x_10; lean_object* x_16; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; lean_object* x_16; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_inc(x_1); x_4 = l_Lake_InputFileConfig_path___proj(x_1); x_5 = lean_ctor_get(x_4, 0); @@ -18185,13 +18185,13 @@ x_11 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldO x_12 = lean_box(0); x_13 = lean_alloc_ctor(3, 1, 1); lean_ctor_set(x_13, 0, x_12); -lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_8); -x_14 = l_Lake_Toml_RBDict_insert___redArg(x_11, x_7, x_13, x_9); +lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_9); +x_14 = l_Lake_Toml_RBDict_insert___redArg(x_11, x_7, x_13, x_8); return x_14; } else { -return x_9; +return x_8; } } block_28: @@ -18221,8 +18221,8 @@ else uint8_t x_24; uint8_t x_25; x_24 = lean_unbox(x_20); x_25 = lean_unbox(x_20); -x_8 = x_24; -x_9 = x_16; +x_8 = x_16; +x_9 = x_24; x_10 = x_25; goto block_15; } @@ -18232,8 +18232,8 @@ else uint8_t x_26; uint8_t x_27; x_26 = lean_unbox(x_20); x_27 = lean_unbox(x_21); -x_8 = x_26; -x_9 = x_16; +x_8 = x_16; +x_9 = x_26; x_10 = x_27; goto block_15; } @@ -19494,7 +19494,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lake_CLI_Translate_Toml_0__Lake_PackageConfig_toToml(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_25; lean_object* x_26; lean_object* x_42; lean_object* x_43; lean_object* x_57; lean_object* x_58; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_88; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_113; lean_object* x_127; lean_object* x_128; lean_object* x_144; lean_object* x_145; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_172; lean_object* x_186; lean_object* x_187; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_212; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_237; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_262; lean_object* x_276; lean_object* x_277; lean_object* x_291; lean_object* x_292; lean_object* x_311; lean_object* x_312; lean_object* x_330; uint8_t x_331; lean_object* x_332; uint8_t x_333; lean_object* x_339; lean_object* x_352; uint8_t x_353; lean_object* x_354; uint8_t x_355; lean_object* x_361; lean_object* x_374; lean_object* x_375; uint8_t x_376; uint8_t x_377; lean_object* x_383; lean_object* x_396; lean_object* x_397; lean_object* x_408; lean_object* x_409; uint8_t x_410; uint8_t x_411; lean_object* x_417; lean_object* x_430; lean_object* x_431; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_457; lean_object* x_468; lean_object* x_469; lean_object* x_481; lean_object* x_482; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_505; lean_object* x_516; lean_object* x_517; lean_object* x_529; lean_object* x_530; lean_object* x_539; lean_object* x_540; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_564; lean_object* x_575; lean_object* x_576; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_599; lean_object* x_610; lean_object* x_611; lean_object* x_623; lean_object* x_624; uint8_t x_625; uint8_t x_626; lean_object* x_632; lean_object* x_645; lean_object* x_646; lean_object* x_656; lean_object* x_657; lean_object* x_667; lean_object* x_668; lean_object* x_683; lean_object* x_684; lean_object* x_699; lean_object* x_700; lean_object* x_715; lean_object* x_716; lean_object* x_731; lean_object* x_732; lean_object* x_747; lean_object* x_748; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_774; lean_object* x_785; uint8_t x_786; lean_object* x_787; uint8_t x_788; lean_object* x_794; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_818; lean_object* x_829; lean_object* x_830; lean_object* x_841; lean_object* x_842; lean_object* x_843; uint8_t x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; uint8_t x_850; lean_object* x_855; uint8_t x_856; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_25; lean_object* x_26; lean_object* x_42; lean_object* x_43; lean_object* x_57; lean_object* x_58; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_88; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_113; lean_object* x_127; lean_object* x_128; lean_object* x_144; lean_object* x_145; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_172; lean_object* x_186; lean_object* x_187; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_212; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_237; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_262; lean_object* x_276; lean_object* x_277; lean_object* x_291; lean_object* x_292; lean_object* x_311; lean_object* x_312; lean_object* x_330; uint8_t x_331; lean_object* x_332; uint8_t x_333; lean_object* x_339; lean_object* x_352; lean_object* x_353; uint8_t x_354; uint8_t x_355; lean_object* x_361; lean_object* x_374; uint8_t x_375; lean_object* x_376; uint8_t x_377; lean_object* x_383; lean_object* x_396; lean_object* x_397; lean_object* x_408; lean_object* x_409; uint8_t x_410; uint8_t x_411; lean_object* x_417; lean_object* x_430; lean_object* x_431; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_457; lean_object* x_468; lean_object* x_469; lean_object* x_481; lean_object* x_482; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_505; lean_object* x_516; lean_object* x_517; lean_object* x_529; lean_object* x_530; lean_object* x_539; lean_object* x_540; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_564; lean_object* x_575; lean_object* x_576; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_599; lean_object* x_610; lean_object* x_611; lean_object* x_623; lean_object* x_624; uint8_t x_625; uint8_t x_626; lean_object* x_632; lean_object* x_645; lean_object* x_646; lean_object* x_656; lean_object* x_657; lean_object* x_667; lean_object* x_668; lean_object* x_683; lean_object* x_684; lean_object* x_699; lean_object* x_700; lean_object* x_715; lean_object* x_716; lean_object* x_731; lean_object* x_732; lean_object* x_747; lean_object* x_748; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_774; lean_object* x_785; uint8_t x_786; lean_object* x_787; uint8_t x_788; lean_object* x_794; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_818; lean_object* x_829; lean_object* x_830; lean_object* x_841; lean_object* x_842; lean_object* x_843; uint8_t x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; uint8_t x_850; lean_object* x_855; uint8_t x_856; x_5 = l_Lake_PackageConfig_bootstrap___proj(x_1, x_2); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -19801,13 +19801,13 @@ goto block_76; lean_object* x_105; lean_object* x_106; size_t x_107; size_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; x_105 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_106 = lean_box(0); -x_107 = lean_array_size(x_103); +x_107 = lean_array_size(x_104); x_108 = 0; -x_109 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_107, x_108, x_103); +x_109 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_107, x_108, x_104); x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_106); lean_ctor_set(x_110, 1, x_109); -x_111 = l_Lake_Toml_RBDict_insert___redArg(x_105, x_102, x_110, x_104); +x_111 = l_Lake_Toml_RBDict_insert___redArg(x_105, x_102, x_110, x_103); x_88 = x_111; goto block_101; } @@ -19836,8 +19836,8 @@ if (x_124 == 0) { lean_dec(x_122); lean_dec_ref(x_121); -x_103 = x_120; -x_104 = x_113; +x_103 = x_113; +x_104 = x_120; goto block_112; } else @@ -19847,8 +19847,8 @@ x_125 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_121); if (x_125 == 0) { -x_103 = x_120; -x_104 = x_113; +x_103 = x_113; +x_104 = x_120; goto block_112; } else @@ -20034,13 +20034,13 @@ goto block_185; lean_object* x_204; lean_object* x_205; size_t x_206; size_t x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; x_204 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_205 = lean_box(0); -x_206 = lean_array_size(x_202); +x_206 = lean_array_size(x_203); x_207 = 0; -x_208 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_206, x_207, x_202); +x_208 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_206, x_207, x_203); x_209 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_209, 0, x_205); lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lake_Toml_RBDict_insert___redArg(x_204, x_201, x_209, x_203); +x_210 = l_Lake_Toml_RBDict_insert___redArg(x_204, x_201, x_209, x_202); x_187 = x_210; goto block_200; } @@ -20069,8 +20069,8 @@ if (x_223 == 0) { lean_dec(x_221); lean_dec_ref(x_220); -x_202 = x_219; -x_203 = x_212; +x_202 = x_212; +x_203 = x_219; goto block_211; } else @@ -20080,8 +20080,8 @@ x_224 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_220); if (x_224 == 0) { -x_202 = x_219; -x_203 = x_212; +x_202 = x_212; +x_203 = x_219; goto block_211; } else @@ -20097,13 +20097,13 @@ goto block_200; lean_object* x_229; lean_object* x_230; size_t x_231; size_t x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; x_229 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_230 = lean_box(0); -x_231 = lean_array_size(x_228); +x_231 = lean_array_size(x_227); x_232 = 0; -x_233 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_231, x_232, x_228); +x_233 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_231, x_232, x_227); 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_Lake_Toml_RBDict_insert___redArg(x_229, x_226, x_234, x_227); +x_235 = l_Lake_Toml_RBDict_insert___redArg(x_229, x_226, x_234, x_228); x_212 = x_235; goto block_225; } @@ -20132,8 +20132,8 @@ if (x_248 == 0) { lean_dec(x_246); lean_dec_ref(x_245); -x_227 = x_237; -x_228 = x_244; +x_227 = x_244; +x_228 = x_237; goto block_236; } else @@ -20143,8 +20143,8 @@ x_249 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_245); if (x_249 == 0) { -x_227 = x_237; -x_228 = x_244; +x_227 = x_244; +x_228 = x_237; goto block_236; } else @@ -20410,14 +20410,14 @@ x_356 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeField x_357 = lean_box(0); x_358 = lean_alloc_ctor(3, 1, 1); lean_ctor_set(x_358, 0, x_357); -lean_ctor_set_uint8(x_358, sizeof(void*)*1, x_353); -x_359 = l_Lake_Toml_RBDict_insert___redArg(x_356, x_352, x_358, x_354); +lean_ctor_set_uint8(x_358, sizeof(void*)*1, x_354); +x_359 = l_Lake_Toml_RBDict_insert___redArg(x_356, x_352, x_358, x_353); x_339 = x_359; goto block_351; } else { -x_339 = x_354; +x_339 = x_353; goto block_351; } } @@ -20449,8 +20449,8 @@ else uint8_t x_369; uint8_t x_370; x_369 = lean_unbox(x_365); x_370 = lean_unbox(x_365); -x_353 = x_369; -x_354 = x_361; +x_353 = x_361; +x_354 = x_369; x_355 = x_370; goto block_360; } @@ -20460,8 +20460,8 @@ else uint8_t x_371; uint8_t x_372; x_371 = lean_unbox(x_365); x_372 = lean_unbox(x_366); -x_353 = x_371; -x_354 = x_361; +x_353 = x_361; +x_354 = x_371; x_355 = x_372; goto block_360; } @@ -20475,14 +20475,14 @@ x_378 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeField x_379 = lean_box(0); x_380 = lean_alloc_ctor(3, 1, 1); lean_ctor_set(x_380, 0, x_379); -lean_ctor_set_uint8(x_380, sizeof(void*)*1, x_376); -x_381 = l_Lake_Toml_RBDict_insert___redArg(x_378, x_374, x_380, x_375); +lean_ctor_set_uint8(x_380, sizeof(void*)*1, x_375); +x_381 = l_Lake_Toml_RBDict_insert___redArg(x_378, x_374, x_380, x_376); x_361 = x_381; goto block_373; } else { -x_361 = x_375; +x_361 = x_376; goto block_373; } } @@ -20514,8 +20514,8 @@ else uint8_t x_391; uint8_t x_392; x_391 = lean_unbox(x_387); x_392 = lean_unbox(x_387); -x_375 = x_383; -x_376 = x_391; +x_375 = x_391; +x_376 = x_383; x_377 = x_392; goto block_382; } @@ -20525,8 +20525,8 @@ else uint8_t x_393; uint8_t x_394; x_393 = lean_unbox(x_387); x_394 = lean_unbox(x_388); -x_375 = x_383; -x_376 = x_393; +x_375 = x_393; +x_376 = x_383; x_377 = x_394; goto block_382; } @@ -20672,13 +20672,13 @@ goto block_429; lean_object* x_449; lean_object* x_450; size_t x_451; size_t x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; x_449 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_450 = lean_box(0); -x_451 = lean_array_size(x_448); +x_451 = lean_array_size(x_447); x_452 = 0; -x_453 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_PackageConfig_toToml_spec__0(x_451, x_452, x_448); +x_453 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_CLI_Translate_Toml_0__Lake_PackageConfig_toToml_spec__0(x_451, x_452, x_447); x_454 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_454, 0, x_450); lean_ctor_set(x_454, 1, x_453); -x_455 = l_Lake_Toml_RBDict_insert___redArg(x_449, x_446, x_454, x_447); +x_455 = l_Lake_Toml_RBDict_insert___redArg(x_449, x_446, x_454, x_448); x_431 = x_455; goto block_445; } @@ -20703,8 +20703,8 @@ if (x_465 == 0) { lean_dec(x_463); lean_dec_ref(x_462); -x_447 = x_457; -x_448 = x_461; +x_447 = x_461; +x_448 = x_457; goto block_456; } else @@ -20714,8 +20714,8 @@ x_466 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Pack lean_dec_ref(x_462); if (x_466 == 0) { -x_447 = x_457; -x_448 = x_461; +x_447 = x_461; +x_448 = x_457; goto block_456; } else @@ -20954,13 +20954,13 @@ goto block_538; lean_object* x_556; lean_object* x_557; size_t x_558; size_t x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; x_556 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_557 = lean_box(0); -x_558 = lean_array_size(x_554); +x_558 = lean_array_size(x_555); x_559 = 0; -x_560 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_558, x_559, x_554); +x_560 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_558, x_559, x_555); x_561 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_561, 0, x_557); lean_ctor_set(x_561, 1, x_560); -x_562 = l_Lake_Toml_RBDict_insert___redArg(x_556, x_553, x_561, x_555); +x_562 = l_Lake_Toml_RBDict_insert___redArg(x_556, x_553, x_561, x_554); x_540 = x_562; goto block_552; } @@ -20985,8 +20985,8 @@ if (x_572 == 0) { lean_dec(x_570); lean_dec_ref(x_569); -x_554 = x_568; -x_555 = x_564; +x_554 = x_564; +x_555 = x_568; goto block_563; } else @@ -20996,8 +20996,8 @@ x_573 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_569); if (x_573 == 0) { -x_554 = x_568; -x_555 = x_564; +x_554 = x_564; +x_555 = x_568; goto block_563; } else @@ -21047,13 +21047,13 @@ goto block_574; lean_object* x_591; lean_object* x_592; size_t x_593; size_t x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; x_591 = l___private_Lake_CLI_Translate_Toml_0__Lake_instInsertFieldOfEncodeFieldOfBEqOfConfigField___redArg___lam__0___closed__0; x_592 = lean_box(0); -x_593 = lean_array_size(x_590); +x_593 = lean_array_size(x_589); x_594 = 0; -x_595 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_593, x_594, x_590); +x_595 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_StrPatDescr_toToml_spec__0(x_593, x_594, x_589); x_596 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_596, 0, x_592); lean_ctor_set(x_596, 1, x_595); -x_597 = l_Lake_Toml_RBDict_insert___redArg(x_591, x_588, x_596, x_589); +x_597 = l_Lake_Toml_RBDict_insert___redArg(x_591, x_588, x_596, x_590); x_576 = x_597; goto block_587; } @@ -21078,8 +21078,8 @@ if (x_607 == 0) { lean_dec(x_605); lean_dec_ref(x_604); -x_589 = x_599; -x_590 = x_603; +x_589 = x_603; +x_590 = x_599; goto block_598; } else @@ -21089,8 +21089,8 @@ x_608 = l_Array_isEqvAux___at___00__private_Lake_CLI_Translate_Toml_0__Lake_Lean lean_dec_ref(x_604); if (x_608 == 0) { -x_589 = x_599; -x_590 = x_603; +x_589 = x_603; +x_590 = x_599; goto block_598; } else diff --git a/stage0/stdlib/Lake/Config/Meta.c b/stage0/stdlib/Lake/Config/Meta.c index de6cb6c867b2..f5bd9561e978 100644 --- a/stage0/stdlib/Lake/Config/Meta.c +++ b/stage0/stdlib/Lake/Config/Meta.c @@ -2411,96 +2411,96 @@ goto block_322; block_74: { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; size_t x_71; size_t x_72; -lean_inc(x_36); -lean_inc(x_35); -x_45 = l_Lean_Syntax_node3(x_35, x_24, x_43, x_36, x_44); -lean_inc_n(x_36, 2); -lean_inc(x_18); -lean_inc(x_35); -x_46 = l_Lean_Syntax_node3(x_35, x_18, x_36, x_36, x_45); -lean_inc(x_35); -x_47 = l_Lean_Syntax_node2(x_35, x_23, x_39, x_46); -lean_inc_n(x_36, 3); -lean_inc(x_18); -lean_inc(x_35); -x_48 = l_Lean_Syntax_node6(x_35, x_18, x_34, x_36, x_21, x_36, x_47, x_36); -lean_inc(x_35); -x_49 = l_Lean_Syntax_node1(x_35, x_30, x_48); +lean_inc(x_22); +lean_inc(x_30); +x_45 = l_Lean_Syntax_node3(x_30, x_40, x_38, x_22, x_44); +lean_inc_n(x_22, 2); +lean_inc(x_25); +lean_inc(x_30); +x_46 = l_Lean_Syntax_node3(x_30, x_25, x_22, x_22, x_45); +lean_inc(x_30); +x_47 = l_Lean_Syntax_node2(x_30, x_26, x_39, x_46); +lean_inc_n(x_22, 3); +lean_inc(x_25); +lean_inc(x_30); +x_48 = l_Lean_Syntax_node6(x_30, x_25, x_34, x_22, x_27, x_22, x_47, x_22); +lean_inc(x_30); +x_49 = l_Lean_Syntax_node1(x_30, x_19, x_48); x_50 = 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__2_spec__2___closed__0; -x_51 = l_Lean_Name_mkStr4(x_42, x_32, x_41, x_50); -lean_inc(x_36); -lean_inc(x_35); -x_52 = l_Lean_Syntax_node1(x_35, x_51, x_36); -lean_inc(x_35); +x_51 = l_Lean_Name_mkStr4(x_18, x_37, x_33, x_50); +lean_inc(x_22); +lean_inc(x_30); +x_52 = l_Lean_Syntax_node1(x_30, x_51, x_22); +lean_inc(x_30); x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_35); -lean_ctor_set(x_53, 1, x_37); +lean_ctor_set(x_53, 0, x_30); +lean_ctor_set(x_53, 1, x_20); x_54 = 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__2_spec__2___closed__1; x_55 = 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__2_spec__2___closed__2; x_56 = 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__2_spec__2___closed__3; -x_57 = l_Lean_addMacroScope(x_22, x_56, x_38); -x_58 = l_Lean_Name_mkStr2(x_26, x_54); +x_57 = l_Lean_addMacroScope(x_24, x_56, x_31); +x_58 = l_Lean_Name_mkStr2(x_28, x_54); lean_inc(x_58); x_59 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_31); +lean_ctor_set(x_59, 1, x_36); x_60 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_60, 0, x_58); x_61 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_19); +lean_ctor_set(x_61, 1, x_29); x_62 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_62, 0, x_59); lean_ctor_set(x_62, 1, x_61); -lean_inc(x_35); +lean_inc(x_30); x_63 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_63, 0, x_35); +lean_ctor_set(x_63, 0, x_30); lean_ctor_set(x_63, 1, x_55); lean_ctor_set(x_63, 2, x_57); lean_ctor_set(x_63, 3, x_62); -lean_inc(x_18); -lean_inc(x_35); -x_64 = l_Lean_Syntax_node2(x_35, x_18, x_53, x_63); +lean_inc(x_25); +lean_inc(x_30); +x_64 = l_Lean_Syntax_node2(x_30, x_25, x_53, x_63); x_65 = 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__2_spec__2___closed__4; -lean_inc(x_35); +lean_inc(x_30); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_35); +lean_ctor_set(x_66, 0, x_30); lean_ctor_set(x_66, 1, x_65); -lean_inc(x_35); -x_67 = l_Lean_Syntax_node6(x_35, x_33, x_40, x_36, x_49, x_52, x_64, x_66); -lean_inc(x_35); -x_68 = l_Lean_Syntax_node1(x_35, x_18, x_67); -x_69 = l_Lean_Syntax_node4(x_35, x_28, x_16, x_20, x_27, x_68); +lean_inc(x_30); +x_67 = l_Lean_Syntax_node6(x_30, x_23, x_42, x_22, x_49, x_52, x_64, x_66); +lean_inc(x_30); +x_68 = l_Lean_Syntax_node1(x_30, x_25, x_67); +x_69 = l_Lean_Syntax_node4(x_30, x_32, x_16, x_41, x_43, x_68); if (lean_is_scalar(x_17)) { x_70 = lean_alloc_ctor(0, 2, 0); } else { x_70 = x_17; } -lean_ctor_set(x_70, 0, x_25); +lean_ctor_set(x_70, 0, x_21); lean_ctor_set(x_70, 1, x_69); x_71 = 1; x_72 = lean_usize_add(x_9, x_71); x_9 = x_72; x_11 = x_70; -x_13 = x_29; +x_13 = x_35; goto _start; } block_299: { lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_96 = l_Lean_replaceRef(x_16, x_90); -lean_dec(x_90); +x_96 = l_Lean_replaceRef(x_16, x_78); +lean_dec(x_78); lean_inc(x_96); -lean_inc(x_85); -lean_inc(x_81); +lean_inc(x_83); +lean_inc(x_87); x_97 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_97, 0, x_91); -lean_ctor_set(x_97, 1, x_81); -lean_ctor_set(x_97, 2, x_85); -lean_ctor_set(x_97, 3, x_84); -lean_ctor_set(x_97, 4, x_89); +lean_ctor_set(x_97, 0, x_88); +lean_ctor_set(x_97, 1, x_87); +lean_ctor_set(x_97, 2, x_83); +lean_ctor_set(x_97, 3, x_93); +lean_ctor_set(x_97, 4, x_92); lean_ctor_set(x_97, 5, x_96); -x_98 = 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__2_spec__2___lam__1(x_14, x_96, x_97, x_86); +x_98 = 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__2_spec__2___lam__1(x_14, x_96, x_97, x_84); lean_dec_ref(x_97); lean_dec(x_96); x_99 = !lean_is_exclusive(x_98); @@ -2509,139 +2509,139 @@ if (x_99 == 0) lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; x_100 = lean_ctor_get(x_98, 0); x_101 = lean_ctor_get(x_98, 1); -lean_inc_ref(x_93); -x_102 = l_Array_append___redArg(x_93, x_95); +lean_inc_ref(x_81); +x_102 = l_Array_append___redArg(x_81, x_95); lean_dec_ref(x_95); -lean_inc(x_80); -lean_inc(x_79); +lean_inc(x_89); +lean_inc(x_86); x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_79); -lean_ctor_set(x_103, 1, x_80); +lean_ctor_set(x_103, 0, x_86); +lean_ctor_set(x_103, 1, x_89); lean_ctor_set(x_103, 2, x_102); -lean_inc_n(x_87, 6); -lean_inc(x_79); -x_104 = l_Lean_Syntax_node7(x_79, x_94, x_87, x_87, x_103, x_87, x_87, x_87, x_87); +lean_inc_n(x_94, 6); +lean_inc(x_86); +x_104 = l_Lean_Syntax_node7(x_86, x_90, x_94, x_94, x_103, x_94, x_94, x_94, x_94); x_105 = 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__2_spec__2___closed__5; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_106 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_105); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_106 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_105); x_107 = l_Lake_configDecl___closed__26; x_108 = 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__2_spec__2___closed__6; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_109 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_108); -lean_inc(x_87); -lean_inc(x_79); -x_110 = l_Lean_Syntax_node1(x_79, x_109, x_87); -lean_inc(x_79); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_109 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_108); +lean_inc(x_94); +lean_inc(x_86); +x_110 = l_Lean_Syntax_node1(x_86, x_109, x_94); +lean_inc(x_86); lean_ctor_set_tag(x_98, 2); lean_ctor_set(x_98, 1, x_105); -lean_ctor_set(x_98, 0, x_79); +lean_ctor_set(x_98, 0, x_86); x_111 = l_Lake_configDecl___closed__8; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_112 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_111); -lean_inc(x_87); -lean_inc(x_79); -x_113 = l_Lean_Syntax_node2(x_79, x_112, x_78, x_87); -lean_inc(x_80); -lean_inc(x_79); -x_114 = l_Lean_Syntax_node1(x_79, x_80, x_113); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_112 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_111); +lean_inc(x_94); +lean_inc(x_86); +x_113 = l_Lean_Syntax_node2(x_86, x_112, x_80, x_94); +lean_inc(x_89); +lean_inc(x_86); +x_114 = l_Lean_Syntax_node1(x_86, x_89, x_113); x_115 = l_Lake_configField___closed__19; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_116 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_115); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_116 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_115); x_117 = 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__2_spec__2___closed__7; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_118 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_117); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_118 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_117); x_119 = 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__2_spec__2___closed__8; -lean_inc(x_79); +lean_inc(x_86); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_79); +lean_ctor_set(x_120, 0, x_86); lean_ctor_set(x_120, 1, x_119); x_121 = 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__2_spec__2___closed__9; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_122 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_121); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_122 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_121); x_123 = 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__2_spec__2___closed__11; x_124 = 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__2_spec__2___closed__12; -lean_inc(x_85); -lean_inc(x_81); -x_125 = l_Lean_addMacroScope(x_81, x_124, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_125 = l_Lean_addMacroScope(x_87, x_124, x_83); x_126 = l_Lake_configField___closed__1; x_127 = lean_box(0); x_128 = 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__2_spec__2___closed__17; -lean_inc(x_79); +lean_inc(x_86); x_129 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_129, 0, x_79); +lean_ctor_set(x_129, 0, x_86); lean_ctor_set(x_129, 1, x_123); lean_ctor_set(x_129, 2, x_125); lean_ctor_set(x_129, 3, x_128); lean_inc(x_3); lean_inc(x_77); lean_inc(x_2); -lean_inc(x_80); -lean_inc(x_79); -x_130 = l_Lean_Syntax_node3(x_79, x_80, x_2, x_77, x_3); -lean_inc(x_79); -x_131 = l_Lean_Syntax_node2(x_79, x_122, x_129, x_130); -lean_inc(x_79); -x_132 = l_Lean_Syntax_node2(x_79, x_118, x_120, x_131); -lean_inc(x_87); -lean_inc(x_79); -x_133 = l_Lean_Syntax_node2(x_79, x_116, x_87, x_132); +lean_inc(x_89); +lean_inc(x_86); +x_130 = l_Lean_Syntax_node3(x_86, x_89, x_2, x_77, x_3); +lean_inc(x_86); +x_131 = l_Lean_Syntax_node2(x_86, x_122, x_129, x_130); +lean_inc(x_86); +x_132 = l_Lean_Syntax_node2(x_86, x_118, x_120, x_131); +lean_inc(x_94); +lean_inc(x_86); +x_133 = l_Lean_Syntax_node2(x_86, x_116, x_94, x_132); x_134 = 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__2_spec__2___closed__18; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_135 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_134); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_135 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_134); x_136 = 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__2_spec__2___closed__19; -lean_inc(x_79); +lean_inc(x_86); x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_79); +lean_ctor_set(x_137, 0, x_86); lean_ctor_set(x_137, 1, 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__2_spec__2___closed__20; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_139 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_138); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_139 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_138); x_140 = 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__2_spec__2___closed__21; -lean_inc(x_79); +lean_inc(x_86); x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_79); +lean_ctor_set(x_141, 0, x_86); lean_ctor_set(x_141, 1, x_140); lean_inc(x_4); -lean_inc(x_80); -lean_inc(x_79); -x_142 = l_Lean_Syntax_node1(x_79, x_80, x_4); +lean_inc(x_89); +lean_inc(x_86); +x_142 = l_Lean_Syntax_node1(x_86, x_89, x_4); x_143 = 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__2_spec__2___closed__22; -lean_inc(x_79); +lean_inc(x_86); x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_79); +lean_ctor_set(x_144, 0, x_86); lean_ctor_set(x_144, 1, x_143); -lean_inc(x_79); -x_145 = l_Lean_Syntax_node3(x_79, x_139, x_141, x_142, x_144); +lean_inc(x_86); +x_145 = l_Lean_Syntax_node3(x_86, x_139, x_141, x_142, x_144); x_146 = 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__2_spec__2___closed__23; x_147 = 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__2_spec__2___closed__24; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_148 = l_Lean_Name_mkStr4(x_92, x_88, x_146, x_147); -lean_inc_n(x_87, 2); -lean_inc(x_79); -x_149 = l_Lean_Syntax_node2(x_79, x_148, x_87, x_87); -lean_inc(x_87); -lean_inc(x_79); -x_150 = l_Lean_Syntax_node4(x_79, x_135, x_137, x_145, x_149, x_87); -lean_inc(x_79); -x_151 = l_Lean_Syntax_node6(x_79, x_106, x_110, x_98, x_87, x_114, x_133, x_150); -x_152 = l_Lean_Syntax_node2(x_79, x_83, x_104, x_151); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_148 = l_Lean_Name_mkStr4(x_85, x_91, x_146, x_147); +lean_inc_n(x_94, 2); +lean_inc(x_86); +x_149 = l_Lean_Syntax_node2(x_86, x_148, x_94, x_94); +lean_inc(x_94); +lean_inc(x_86); +x_150 = l_Lean_Syntax_node4(x_86, x_135, x_137, x_145, x_149, x_94); +lean_inc(x_86); +x_151 = l_Lean_Syntax_node6(x_86, x_106, x_110, x_98, x_94, x_114, x_133, x_150); +x_152 = l_Lean_Syntax_node2(x_86, x_79, x_104, x_151); x_153 = lean_array_push(x_15, x_152); x_154 = 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__2_spec__2___closed__25; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_155 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_154); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_155 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_154); x_156 = 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__2_spec__2___closed__26; lean_inc(x_100); x_157 = lean_alloc_ctor(2, 2, 0); @@ -2649,9 +2649,9 @@ lean_ctor_set(x_157, 0, x_100); lean_ctor_set(x_157, 1, x_156); x_158 = 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__2_spec__2___closed__28; x_159 = 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__2_spec__2___closed__29; -lean_inc(x_85); -lean_inc(x_81); -x_160 = l_Lean_addMacroScope(x_81, x_159, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_160 = l_Lean_addMacroScope(x_87, x_159, x_83); lean_inc(x_100); x_161 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_161, 0, x_100); @@ -2659,37 +2659,37 @@ lean_ctor_set(x_161, 1, x_158); lean_ctor_set(x_161, 2, x_160); lean_ctor_set(x_161, 3, x_127); x_162 = 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__2_spec__2___closed__30; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_163 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_162); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_163 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_162); x_164 = 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__2_spec__2___closed__31; lean_inc(x_100); x_165 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_165, 0, x_100); lean_ctor_set(x_165, 1, x_164); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_100); x_166 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_166, 0, x_100); -lean_ctor_set(x_166, 1, x_80); -lean_ctor_set(x_166, 2, x_93); +lean_ctor_set(x_166, 1, x_89); +lean_ctor_set(x_166, 2, x_81); x_167 = 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__2_spec__2___closed__32; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_168 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_167); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_168 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_167); x_169 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_170 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_169); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_170 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_169); x_171 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_172 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_171); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_172 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_171); x_173 = 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__2_spec__2___closed__34; x_174 = 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__2_spec__2___closed__35; -lean_inc(x_85); -lean_inc(x_81); -x_175 = l_Lean_addMacroScope(x_81, x_174, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_175 = l_Lean_addMacroScope(x_87, x_174, x_83); lean_inc(x_100); x_176 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_176, 0, x_100); @@ -2701,9 +2701,9 @@ lean_inc(x_172); lean_inc(x_100); x_177 = l_Lean_Syntax_node2(x_100, x_172, x_176, x_166); x_178 = 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__2_spec__2___closed__36; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_179 = l_Lean_Name_mkStr4(x_92, x_88, x_107, x_178); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_179 = l_Lean_Name_mkStr4(x_85, x_91, x_107, x_178); lean_inc(x_100); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_100); @@ -2714,17 +2714,17 @@ lean_inc(x_179); lean_inc(x_100); x_181 = l_Lean_Syntax_node3(x_100, x_179, x_180, x_166, x_77); lean_inc_ref_n(x_166, 2); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_100); -x_182 = l_Lean_Syntax_node3(x_100, x_80, x_166, x_166, x_181); +x_182 = l_Lean_Syntax_node3(x_100, x_89, x_166, x_166, x_181); lean_inc(x_170); lean_inc(x_100); x_183 = l_Lean_Syntax_node2(x_100, x_170, x_177, x_182); x_184 = 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__2_spec__2___closed__38; x_185 = 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__2_spec__2___closed__39; -lean_inc(x_85); -lean_inc(x_81); -x_186 = l_Lean_addMacroScope(x_81, x_185, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_186 = l_Lean_addMacroScope(x_87, x_185, x_83); lean_inc(x_100); x_187 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_187, 0, x_100); @@ -2742,17 +2742,17 @@ lean_inc(x_179); lean_inc(x_100); x_189 = l_Lean_Syntax_node3(x_100, x_179, x_180, x_166, x_5); lean_inc_ref_n(x_166, 2); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_100); -x_190 = l_Lean_Syntax_node3(x_100, x_80, x_166, x_166, x_189); +x_190 = l_Lean_Syntax_node3(x_100, x_89, x_166, x_166, x_189); lean_inc(x_170); lean_inc(x_100); x_191 = l_Lean_Syntax_node2(x_100, x_170, x_188, x_190); x_192 = 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__2_spec__2___closed__41; x_193 = 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__2_spec__2___closed__42; -lean_inc(x_85); -lean_inc(x_81); -x_194 = l_Lean_addMacroScope(x_81, x_193, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_194 = l_Lean_addMacroScope(x_87, x_193, x_83); lean_inc(x_100); x_195 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_195, 0, x_100); @@ -2766,32 +2766,32 @@ if (x_1 == 0) { lean_object* x_197; x_197 = 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__2_spec__2___closed__46; -x_18 = x_80; -x_19 = x_127; -x_20 = x_157; -x_21 = x_191; -x_22 = x_81; -x_23 = x_170; -x_24 = x_179; -x_25 = x_153; -x_26 = x_126; -x_27 = x_161; -x_28 = x_155; -x_29 = x_101; -x_30 = x_168; -x_31 = x_127; -x_32 = x_88; -x_33 = x_163; +x_18 = x_85; +x_19 = x_168; +x_20 = x_119; +x_21 = x_153; +x_22 = x_166; +x_23 = x_163; +x_24 = x_87; +x_25 = x_89; +x_26 = x_170; +x_27 = x_191; +x_28 = x_126; +x_29 = x_127; +x_30 = x_100; +x_31 = x_83; +x_32 = x_155; +x_33 = x_107; x_34 = x_183; -x_35 = x_100; -x_36 = x_166; -x_37 = x_119; -x_38 = x_85; +x_35 = x_101; +x_36 = x_127; +x_37 = x_91; +x_38 = x_180; x_39 = x_196; -x_40 = x_165; -x_41 = x_107; -x_42 = x_92; -x_43 = x_180; +x_40 = x_179; +x_41 = x_157; +x_42 = x_165; +x_43 = x_161; x_44 = x_197; goto block_74; } @@ -2799,32 +2799,32 @@ else { lean_object* x_198; x_198 = 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__2_spec__2___closed__49; -x_18 = x_80; -x_19 = x_127; -x_20 = x_157; -x_21 = x_191; -x_22 = x_81; -x_23 = x_170; -x_24 = x_179; -x_25 = x_153; -x_26 = x_126; -x_27 = x_161; -x_28 = x_155; -x_29 = x_101; -x_30 = x_168; -x_31 = x_127; -x_32 = x_88; -x_33 = x_163; +x_18 = x_85; +x_19 = x_168; +x_20 = x_119; +x_21 = x_153; +x_22 = x_166; +x_23 = x_163; +x_24 = x_87; +x_25 = x_89; +x_26 = x_170; +x_27 = x_191; +x_28 = x_126; +x_29 = x_127; +x_30 = x_100; +x_31 = x_83; +x_32 = x_155; +x_33 = x_107; x_34 = x_183; -x_35 = x_100; -x_36 = x_166; -x_37 = x_119; -x_38 = x_85; +x_35 = x_101; +x_36 = x_127; +x_37 = x_91; +x_38 = x_180; x_39 = x_196; -x_40 = x_165; -x_41 = x_107; -x_42 = x_92; -x_43 = x_180; +x_40 = x_179; +x_41 = x_157; +x_42 = x_165; +x_43 = x_161; x_44 = x_198; goto block_74; } @@ -2837,139 +2837,139 @@ x_200 = lean_ctor_get(x_98, 1); lean_inc(x_200); lean_inc(x_199); lean_dec(x_98); -lean_inc_ref(x_93); -x_201 = l_Array_append___redArg(x_93, x_95); +lean_inc_ref(x_81); +x_201 = l_Array_append___redArg(x_81, x_95); lean_dec_ref(x_95); -lean_inc(x_80); -lean_inc(x_79); +lean_inc(x_89); +lean_inc(x_86); x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_79); -lean_ctor_set(x_202, 1, x_80); +lean_ctor_set(x_202, 0, x_86); +lean_ctor_set(x_202, 1, x_89); lean_ctor_set(x_202, 2, x_201); -lean_inc_n(x_87, 6); -lean_inc(x_79); -x_203 = l_Lean_Syntax_node7(x_79, x_94, x_87, x_87, x_202, x_87, x_87, x_87, x_87); +lean_inc_n(x_94, 6); +lean_inc(x_86); +x_203 = l_Lean_Syntax_node7(x_86, x_90, x_94, x_94, x_202, x_94, x_94, x_94, x_94); x_204 = 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__2_spec__2___closed__5; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_205 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_204); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_205 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_204); x_206 = l_Lake_configDecl___closed__26; x_207 = 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__2_spec__2___closed__6; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_208 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_207); -lean_inc(x_87); -lean_inc(x_79); -x_209 = l_Lean_Syntax_node1(x_79, x_208, x_87); -lean_inc(x_79); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_208 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_207); +lean_inc(x_94); +lean_inc(x_86); +x_209 = l_Lean_Syntax_node1(x_86, x_208, x_94); +lean_inc(x_86); x_210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_210, 0, x_79); +lean_ctor_set(x_210, 0, x_86); lean_ctor_set(x_210, 1, x_204); x_211 = l_Lake_configDecl___closed__8; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_212 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_211); -lean_inc(x_87); -lean_inc(x_79); -x_213 = l_Lean_Syntax_node2(x_79, x_212, x_78, x_87); -lean_inc(x_80); -lean_inc(x_79); -x_214 = l_Lean_Syntax_node1(x_79, x_80, x_213); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_212 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_211); +lean_inc(x_94); +lean_inc(x_86); +x_213 = l_Lean_Syntax_node2(x_86, x_212, x_80, x_94); +lean_inc(x_89); +lean_inc(x_86); +x_214 = l_Lean_Syntax_node1(x_86, x_89, x_213); x_215 = l_Lake_configField___closed__19; lean_inc_ref(x_82); -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_216 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_215); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_216 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_215); x_217 = 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__2_spec__2___closed__7; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_218 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_217); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_218 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_217); x_219 = 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__2_spec__2___closed__8; -lean_inc(x_79); +lean_inc(x_86); x_220 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_220, 0, x_79); +lean_ctor_set(x_220, 0, x_86); lean_ctor_set(x_220, 1, x_219); x_221 = 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__2_spec__2___closed__9; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_222 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_221); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_222 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_221); x_223 = 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__2_spec__2___closed__11; x_224 = 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__2_spec__2___closed__12; -lean_inc(x_85); -lean_inc(x_81); -x_225 = l_Lean_addMacroScope(x_81, x_224, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_225 = l_Lean_addMacroScope(x_87, x_224, x_83); x_226 = l_Lake_configField___closed__1; x_227 = lean_box(0); x_228 = 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__2_spec__2___closed__17; -lean_inc(x_79); +lean_inc(x_86); x_229 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_229, 0, x_79); +lean_ctor_set(x_229, 0, x_86); lean_ctor_set(x_229, 1, x_223); lean_ctor_set(x_229, 2, x_225); lean_ctor_set(x_229, 3, x_228); lean_inc(x_3); lean_inc(x_77); lean_inc(x_2); -lean_inc(x_80); -lean_inc(x_79); -x_230 = l_Lean_Syntax_node3(x_79, x_80, x_2, x_77, x_3); -lean_inc(x_79); -x_231 = l_Lean_Syntax_node2(x_79, x_222, x_229, x_230); -lean_inc(x_79); -x_232 = l_Lean_Syntax_node2(x_79, x_218, x_220, x_231); -lean_inc(x_87); -lean_inc(x_79); -x_233 = l_Lean_Syntax_node2(x_79, x_216, x_87, x_232); +lean_inc(x_89); +lean_inc(x_86); +x_230 = l_Lean_Syntax_node3(x_86, x_89, x_2, x_77, x_3); +lean_inc(x_86); +x_231 = l_Lean_Syntax_node2(x_86, x_222, x_229, x_230); +lean_inc(x_86); +x_232 = l_Lean_Syntax_node2(x_86, x_218, x_220, x_231); +lean_inc(x_94); +lean_inc(x_86); +x_233 = l_Lean_Syntax_node2(x_86, x_216, x_94, x_232); x_234 = 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__2_spec__2___closed__18; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_235 = l_Lean_Name_mkStr4(x_92, x_88, x_82, x_234); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_235 = l_Lean_Name_mkStr4(x_85, x_91, x_82, x_234); x_236 = 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__2_spec__2___closed__19; -lean_inc(x_79); +lean_inc(x_86); x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_79); +lean_ctor_set(x_237, 0, x_86); lean_ctor_set(x_237, 1, x_236); x_238 = 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__2_spec__2___closed__20; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_239 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_238); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_239 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_238); x_240 = 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__2_spec__2___closed__21; -lean_inc(x_79); +lean_inc(x_86); x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_79); +lean_ctor_set(x_241, 0, x_86); lean_ctor_set(x_241, 1, x_240); lean_inc(x_4); -lean_inc(x_80); -lean_inc(x_79); -x_242 = l_Lean_Syntax_node1(x_79, x_80, x_4); +lean_inc(x_89); +lean_inc(x_86); +x_242 = l_Lean_Syntax_node1(x_86, x_89, x_4); x_243 = 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__2_spec__2___closed__22; -lean_inc(x_79); +lean_inc(x_86); x_244 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_244, 0, x_79); +lean_ctor_set(x_244, 0, x_86); lean_ctor_set(x_244, 1, x_243); -lean_inc(x_79); -x_245 = l_Lean_Syntax_node3(x_79, x_239, x_241, x_242, x_244); +lean_inc(x_86); +x_245 = l_Lean_Syntax_node3(x_86, x_239, x_241, x_242, x_244); x_246 = 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__2_spec__2___closed__23; x_247 = 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__2_spec__2___closed__24; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_248 = l_Lean_Name_mkStr4(x_92, x_88, x_246, x_247); -lean_inc_n(x_87, 2); -lean_inc(x_79); -x_249 = l_Lean_Syntax_node2(x_79, x_248, x_87, x_87); -lean_inc(x_87); -lean_inc(x_79); -x_250 = l_Lean_Syntax_node4(x_79, x_235, x_237, x_245, x_249, x_87); -lean_inc(x_79); -x_251 = l_Lean_Syntax_node6(x_79, x_205, x_209, x_210, x_87, x_214, x_233, x_250); -x_252 = l_Lean_Syntax_node2(x_79, x_83, x_203, x_251); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_248 = l_Lean_Name_mkStr4(x_85, x_91, x_246, x_247); +lean_inc_n(x_94, 2); +lean_inc(x_86); +x_249 = l_Lean_Syntax_node2(x_86, x_248, x_94, x_94); +lean_inc(x_94); +lean_inc(x_86); +x_250 = l_Lean_Syntax_node4(x_86, x_235, x_237, x_245, x_249, x_94); +lean_inc(x_86); +x_251 = l_Lean_Syntax_node6(x_86, x_205, x_209, x_210, x_94, x_214, x_233, x_250); +x_252 = l_Lean_Syntax_node2(x_86, x_79, x_203, x_251); x_253 = lean_array_push(x_15, x_252); x_254 = 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__2_spec__2___closed__25; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_255 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_254); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_255 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_254); x_256 = 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__2_spec__2___closed__26; lean_inc(x_199); x_257 = lean_alloc_ctor(2, 2, 0); @@ -2977,9 +2977,9 @@ lean_ctor_set(x_257, 0, x_199); lean_ctor_set(x_257, 1, x_256); x_258 = 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__2_spec__2___closed__28; x_259 = 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__2_spec__2___closed__29; -lean_inc(x_85); -lean_inc(x_81); -x_260 = l_Lean_addMacroScope(x_81, x_259, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_260 = l_Lean_addMacroScope(x_87, x_259, x_83); lean_inc(x_199); x_261 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_261, 0, x_199); @@ -2987,37 +2987,37 @@ lean_ctor_set(x_261, 1, x_258); lean_ctor_set(x_261, 2, x_260); lean_ctor_set(x_261, 3, x_227); x_262 = 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__2_spec__2___closed__30; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_263 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_262); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_263 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_262); x_264 = 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__2_spec__2___closed__31; lean_inc(x_199); x_265 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_265, 0, x_199); lean_ctor_set(x_265, 1, x_264); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_199); x_266 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_266, 0, x_199); -lean_ctor_set(x_266, 1, x_80); -lean_ctor_set(x_266, 2, x_93); +lean_ctor_set(x_266, 1, x_89); +lean_ctor_set(x_266, 2, x_81); x_267 = 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__2_spec__2___closed__32; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_268 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_267); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_268 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_267); x_269 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_270 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_269); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_270 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_269); x_271 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_272 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_271); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_272 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_271); x_273 = 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__2_spec__2___closed__34; x_274 = 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__2_spec__2___closed__35; -lean_inc(x_85); -lean_inc(x_81); -x_275 = l_Lean_addMacroScope(x_81, x_274, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_275 = l_Lean_addMacroScope(x_87, x_274, x_83); lean_inc(x_199); x_276 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_276, 0, x_199); @@ -3029,9 +3029,9 @@ lean_inc(x_272); lean_inc(x_199); x_277 = l_Lean_Syntax_node2(x_199, x_272, x_276, x_266); x_278 = 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__2_spec__2___closed__36; -lean_inc_ref(x_88); -lean_inc_ref(x_92); -x_279 = l_Lean_Name_mkStr4(x_92, x_88, x_206, x_278); +lean_inc_ref(x_91); +lean_inc_ref(x_85); +x_279 = l_Lean_Name_mkStr4(x_85, x_91, x_206, x_278); lean_inc(x_199); x_280 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_280, 0, x_199); @@ -3042,17 +3042,17 @@ lean_inc(x_279); lean_inc(x_199); x_281 = l_Lean_Syntax_node3(x_199, x_279, x_280, x_266, x_77); lean_inc_ref_n(x_266, 2); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_199); -x_282 = l_Lean_Syntax_node3(x_199, x_80, x_266, x_266, x_281); +x_282 = l_Lean_Syntax_node3(x_199, x_89, x_266, x_266, x_281); lean_inc(x_270); lean_inc(x_199); x_283 = l_Lean_Syntax_node2(x_199, x_270, x_277, x_282); x_284 = 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__2_spec__2___closed__38; x_285 = 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__2_spec__2___closed__39; -lean_inc(x_85); -lean_inc(x_81); -x_286 = l_Lean_addMacroScope(x_81, x_285, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_286 = l_Lean_addMacroScope(x_87, x_285, x_83); lean_inc(x_199); x_287 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_287, 0, x_199); @@ -3070,17 +3070,17 @@ lean_inc(x_279); lean_inc(x_199); x_289 = l_Lean_Syntax_node3(x_199, x_279, x_280, x_266, x_5); lean_inc_ref_n(x_266, 2); -lean_inc(x_80); +lean_inc(x_89); lean_inc(x_199); -x_290 = l_Lean_Syntax_node3(x_199, x_80, x_266, x_266, x_289); +x_290 = l_Lean_Syntax_node3(x_199, x_89, x_266, x_266, x_289); lean_inc(x_270); lean_inc(x_199); x_291 = l_Lean_Syntax_node2(x_199, x_270, x_288, x_290); x_292 = 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__2_spec__2___closed__41; x_293 = 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__2_spec__2___closed__42; -lean_inc(x_85); -lean_inc(x_81); -x_294 = l_Lean_addMacroScope(x_81, x_293, x_85); +lean_inc(x_83); +lean_inc(x_87); +x_294 = l_Lean_addMacroScope(x_87, x_293, x_83); lean_inc(x_199); x_295 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_295, 0, x_199); @@ -3094,32 +3094,32 @@ if (x_1 == 0) { lean_object* x_297; x_297 = 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__2_spec__2___closed__46; -x_18 = x_80; -x_19 = x_227; -x_20 = x_257; -x_21 = x_291; -x_22 = x_81; -x_23 = x_270; -x_24 = x_279; -x_25 = x_253; -x_26 = x_226; -x_27 = x_261; -x_28 = x_255; -x_29 = x_200; -x_30 = x_268; -x_31 = x_227; -x_32 = x_88; -x_33 = x_263; +x_18 = x_85; +x_19 = x_268; +x_20 = x_219; +x_21 = x_253; +x_22 = x_266; +x_23 = x_263; +x_24 = x_87; +x_25 = x_89; +x_26 = x_270; +x_27 = x_291; +x_28 = x_226; +x_29 = x_227; +x_30 = x_199; +x_31 = x_83; +x_32 = x_255; +x_33 = x_206; x_34 = x_283; -x_35 = x_199; -x_36 = x_266; -x_37 = x_219; -x_38 = x_85; +x_35 = x_200; +x_36 = x_227; +x_37 = x_91; +x_38 = x_280; x_39 = x_296; -x_40 = x_265; -x_41 = x_206; -x_42 = x_92; -x_43 = x_280; +x_40 = x_279; +x_41 = x_257; +x_42 = x_265; +x_43 = x_261; x_44 = x_297; goto block_74; } @@ -3127,32 +3127,32 @@ else { lean_object* x_298; x_298 = 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__2_spec__2___closed__49; -x_18 = x_80; -x_19 = x_227; -x_20 = x_257; -x_21 = x_291; -x_22 = x_81; -x_23 = x_270; -x_24 = x_279; -x_25 = x_253; -x_26 = x_226; -x_27 = x_261; -x_28 = x_255; -x_29 = x_200; -x_30 = x_268; -x_31 = x_227; -x_32 = x_88; -x_33 = x_263; +x_18 = x_85; +x_19 = x_268; +x_20 = x_219; +x_21 = x_253; +x_22 = x_266; +x_23 = x_263; +x_24 = x_87; +x_25 = x_89; +x_26 = x_270; +x_27 = x_291; +x_28 = x_226; +x_29 = x_227; +x_30 = x_199; +x_31 = x_83; +x_32 = x_255; +x_33 = x_206; x_34 = x_283; -x_35 = x_199; -x_36 = x_266; -x_37 = x_219; -x_38 = x_85; +x_35 = x_200; +x_36 = x_227; +x_37 = x_91; +x_38 = x_280; x_39 = x_296; -x_40 = x_265; -x_41 = x_206; -x_42 = x_92; -x_43 = x_280; +x_40 = x_279; +x_41 = x_257; +x_42 = x_265; +x_43 = x_261; x_44 = x_298; goto block_74; } @@ -3191,29 +3191,29 @@ if (lean_obj_tag(x_6) == 0) { lean_object* x_319; x_319 = 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__2_spec__2___closed__53; -lean_inc(x_301); -lean_inc(x_306); -lean_inc(x_305); -lean_inc(x_303); lean_inc(x_304); +lean_inc(x_305); +lean_inc(x_301); lean_inc(x_302); -x_78 = x_310; -x_79 = x_308; -x_80 = x_316; -x_81 = x_302; +lean_inc(x_303); +lean_inc(x_306); +x_78 = x_306; +x_79 = x_314; +x_80 = x_310; +x_81 = x_317; x_82 = x_313; -x_83 = x_314; -x_84 = x_304; -x_85 = x_303; -x_86 = x_309; -x_87 = x_318; -x_88 = x_312; -x_89 = x_305; -x_90 = x_306; -x_91 = x_301; -x_92 = x_311; -x_93 = x_317; -x_94 = x_315; +x_83 = x_303; +x_84 = x_309; +x_85 = x_311; +x_86 = x_308; +x_87 = x_302; +x_88 = x_301; +x_89 = x_316; +x_90 = x_315; +x_91 = x_312; +x_92 = x_305; +x_93 = x_304; +x_94 = x_318; x_95 = x_319; goto block_299; } @@ -3223,29 +3223,29 @@ lean_object* x_320; lean_object* x_321; x_320 = lean_ctor_get(x_6, 0); lean_inc(x_320); x_321 = l_Array_mkArray1___redArg(x_320); -lean_inc(x_301); -lean_inc(x_306); -lean_inc(x_305); -lean_inc(x_303); lean_inc(x_304); +lean_inc(x_305); +lean_inc(x_301); lean_inc(x_302); -x_78 = x_310; -x_79 = x_308; -x_80 = x_316; -x_81 = x_302; +lean_inc(x_303); +lean_inc(x_306); +x_78 = x_306; +x_79 = x_314; +x_80 = x_310; +x_81 = x_317; x_82 = x_313; -x_83 = x_314; -x_84 = x_304; -x_85 = x_303; -x_86 = x_309; -x_87 = x_318; -x_88 = x_312; -x_89 = x_305; -x_90 = x_306; -x_91 = x_301; -x_92 = x_311; -x_93 = x_317; -x_94 = x_315; +x_83 = x_303; +x_84 = x_309; +x_85 = x_311; +x_86 = x_308; +x_87 = x_302; +x_88 = x_301; +x_89 = x_316; +x_90 = x_315; +x_91 = x_312; +x_92 = x_305; +x_93 = x_304; +x_94 = x_318; x_95 = x_321; goto block_299; } @@ -3341,94 +3341,94 @@ goto block_322; block_74: { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; size_t x_71; size_t x_72; lean_object* x_73; -lean_inc(x_19); lean_inc(x_33); -x_45 = l_Lean_Syntax_node3(x_33, x_23, x_20, x_19, x_44); -lean_inc_n(x_19, 2); lean_inc(x_18); -lean_inc(x_33); -x_46 = l_Lean_Syntax_node3(x_33, x_18, x_19, x_19, x_45); -lean_inc(x_33); -x_47 = l_Lean_Syntax_node2(x_33, x_32, x_34, x_46); -lean_inc_n(x_19, 3); +x_45 = l_Lean_Syntax_node3(x_18, x_25, x_42, x_33, x_44); +lean_inc_n(x_33, 2); +lean_inc(x_31); lean_inc(x_18); -lean_inc(x_33); -x_48 = l_Lean_Syntax_node6(x_33, x_18, x_28, x_19, x_27, x_19, x_47, x_19); -lean_inc(x_33); -x_49 = l_Lean_Syntax_node1(x_33, x_25, x_48); +x_46 = l_Lean_Syntax_node3(x_18, x_31, x_33, x_33, x_45); +lean_inc(x_18); +x_47 = l_Lean_Syntax_node2(x_18, x_37, x_35, x_46); +lean_inc_n(x_33, 3); +lean_inc(x_31); +lean_inc(x_18); +x_48 = l_Lean_Syntax_node6(x_18, x_31, x_21, x_33, x_34, x_33, x_47, x_33); +lean_inc(x_18); +x_49 = l_Lean_Syntax_node1(x_18, x_23, x_48); x_50 = 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__2_spec__2___closed__0; -x_51 = l_Lean_Name_mkStr4(x_31, x_30, x_39, x_50); -lean_inc(x_19); -lean_inc(x_33); -x_52 = l_Lean_Syntax_node1(x_33, x_51, x_19); +x_51 = l_Lean_Name_mkStr4(x_32, x_24, x_39, x_50); lean_inc(x_33); +lean_inc(x_18); +x_52 = l_Lean_Syntax_node1(x_18, x_51, x_33); +lean_inc(x_18); x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_33); -lean_ctor_set(x_53, 1, x_35); +lean_ctor_set(x_53, 0, x_18); +lean_ctor_set(x_53, 1, x_28); x_54 = 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__2_spec__2___closed__1; x_55 = 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__2_spec__2___closed__2; x_56 = 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__2_spec__2___closed__3; -x_57 = l_Lean_addMacroScope(x_26, x_56, x_29); -x_58 = l_Lean_Name_mkStr2(x_41, x_54); +x_57 = l_Lean_addMacroScope(x_43, x_56, x_30); +x_58 = l_Lean_Name_mkStr2(x_40, x_54); lean_inc(x_58); x_59 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_40); +lean_ctor_set(x_59, 1, x_27); x_60 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_60, 0, x_58); x_61 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_21); +lean_ctor_set(x_61, 1, x_36); x_62 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_62, 0, x_59); lean_ctor_set(x_62, 1, x_61); -lean_inc(x_33); +lean_inc(x_18); x_63 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_63, 0, x_33); +lean_ctor_set(x_63, 0, x_18); lean_ctor_set(x_63, 1, x_55); lean_ctor_set(x_63, 2, x_57); lean_ctor_set(x_63, 3, x_62); +lean_inc(x_31); lean_inc(x_18); -lean_inc(x_33); -x_64 = l_Lean_Syntax_node2(x_33, x_18, x_53, x_63); +x_64 = l_Lean_Syntax_node2(x_18, x_31, x_53, x_63); x_65 = 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__2_spec__2___closed__4; -lean_inc(x_33); +lean_inc(x_18); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_33); +lean_ctor_set(x_66, 0, x_18); lean_ctor_set(x_66, 1, x_65); -lean_inc(x_33); -x_67 = l_Lean_Syntax_node6(x_33, x_42, x_22, x_19, x_49, x_52, x_64, x_66); -lean_inc(x_33); -x_68 = l_Lean_Syntax_node1(x_33, x_18, x_67); -x_69 = l_Lean_Syntax_node4(x_33, x_43, x_16, x_24, x_37, x_68); +lean_inc(x_18); +x_67 = l_Lean_Syntax_node6(x_18, x_22, x_38, x_33, x_49, x_52, x_64, x_66); +lean_inc(x_18); +x_68 = l_Lean_Syntax_node1(x_18, x_31, x_67); +x_69 = l_Lean_Syntax_node4(x_18, x_19, x_16, x_20, x_29, x_68); if (lean_is_scalar(x_17)) { x_70 = lean_alloc_ctor(0, 2, 0); } else { x_70 = x_17; } -lean_ctor_set(x_70, 0, x_36); +lean_ctor_set(x_70, 0, x_26); lean_ctor_set(x_70, 1, x_69); x_71 = 1; x_72 = lean_usize_add(x_9, x_71); -x_73 = 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__2_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_72, x_10, x_70, x_12, x_38); +x_73 = 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__2_spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_72, x_10, x_70, x_12, x_41); return x_73; } block_299: { lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_96 = l_Lean_replaceRef(x_16, x_86); -lean_dec(x_86); +x_96 = l_Lean_replaceRef(x_16, x_82); +lean_dec(x_82); lean_inc(x_96); -lean_inc(x_89); -lean_inc(x_88); +lean_inc(x_78); +lean_inc(x_94); x_97 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_97, 0, x_91); -lean_ctor_set(x_97, 1, x_88); -lean_ctor_set(x_97, 2, x_89); -lean_ctor_set(x_97, 3, x_82); -lean_ctor_set(x_97, 4, x_78); +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_94); +lean_ctor_set(x_97, 2, x_78); +lean_ctor_set(x_97, 3, x_83); +lean_ctor_set(x_97, 4, x_90); lean_ctor_set(x_97, 5, x_96); -x_98 = 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__2_spec__2___lam__1(x_14, x_96, x_97, x_84); +x_98 = 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__2_spec__2___lam__1(x_14, x_96, x_97, x_81); lean_dec_ref(x_97); lean_dec(x_96); x_99 = !lean_is_exclusive(x_98); @@ -3441,71 +3441,71 @@ lean_inc_ref(x_87); x_102 = l_Array_append___redArg(x_87, x_95); lean_dec_ref(x_95); lean_inc(x_79); -lean_inc(x_83); +lean_inc(x_84); x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_83); +lean_ctor_set(x_103, 0, x_84); lean_ctor_set(x_103, 1, x_79); lean_ctor_set(x_103, 2, x_102); -lean_inc_n(x_80, 6); -lean_inc(x_83); -x_104 = l_Lean_Syntax_node7(x_83, x_94, x_80, x_80, x_103, x_80, x_80, x_80, x_80); +lean_inc_n(x_85, 6); +lean_inc(x_84); +x_104 = l_Lean_Syntax_node7(x_84, x_93, x_85, x_85, x_103, x_85, x_85, x_85, x_85); x_105 = 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__2_spec__2___closed__5; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_106 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_105); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_106 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_105); x_107 = l_Lake_configDecl___closed__26; x_108 = 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__2_spec__2___closed__6; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_109 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_108); -lean_inc(x_80); -lean_inc(x_83); -x_110 = l_Lean_Syntax_node1(x_83, x_109, x_80); -lean_inc(x_83); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_109 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_108); +lean_inc(x_85); +lean_inc(x_84); +x_110 = l_Lean_Syntax_node1(x_84, x_109, x_85); +lean_inc(x_84); lean_ctor_set_tag(x_98, 2); lean_ctor_set(x_98, 1, x_105); -lean_ctor_set(x_98, 0, x_83); +lean_ctor_set(x_98, 0, x_84); x_111 = l_Lake_configDecl___closed__8; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_112 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_111); -lean_inc(x_80); -lean_inc(x_83); -x_113 = l_Lean_Syntax_node2(x_83, x_112, x_81, x_80); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_112 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_111); +lean_inc(x_85); +lean_inc(x_84); +x_113 = l_Lean_Syntax_node2(x_84, x_112, x_86, x_85); lean_inc(x_79); -lean_inc(x_83); -x_114 = l_Lean_Syntax_node1(x_83, x_79, x_113); +lean_inc(x_84); +x_114 = l_Lean_Syntax_node1(x_84, x_79, x_113); x_115 = l_Lake_configField___closed__19; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_116 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_115); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_116 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_115); x_117 = 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__2_spec__2___closed__7; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_118 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_117); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_118 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_117); x_119 = 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__2_spec__2___closed__8; -lean_inc(x_83); +lean_inc(x_84); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_83); +lean_ctor_set(x_120, 0, x_84); lean_ctor_set(x_120, 1, x_119); x_121 = 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__2_spec__2___closed__9; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_122 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_121); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_122 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_121); x_123 = 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__2_spec__2___closed__11; x_124 = 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__2_spec__2___closed__12; -lean_inc(x_89); -lean_inc(x_88); -x_125 = l_Lean_addMacroScope(x_88, x_124, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_125 = l_Lean_addMacroScope(x_94, x_124, x_78); x_126 = l_Lake_configField___closed__1; x_127 = lean_box(0); x_128 = 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__2_spec__2___closed__17; -lean_inc(x_83); +lean_inc(x_84); x_129 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_129, 0, x_83); +lean_ctor_set(x_129, 0, x_84); lean_ctor_set(x_129, 1, x_123); lean_ctor_set(x_129, 2, x_125); lean_ctor_set(x_129, 3, x_128); @@ -3513,63 +3513,63 @@ lean_inc(x_3); lean_inc(x_77); lean_inc(x_2); lean_inc(x_79); -lean_inc(x_83); -x_130 = l_Lean_Syntax_node3(x_83, x_79, x_2, x_77, x_3); -lean_inc(x_83); -x_131 = l_Lean_Syntax_node2(x_83, x_122, x_129, x_130); -lean_inc(x_83); -x_132 = l_Lean_Syntax_node2(x_83, x_118, x_120, x_131); -lean_inc(x_80); -lean_inc(x_83); -x_133 = l_Lean_Syntax_node2(x_83, x_116, x_80, x_132); +lean_inc(x_84); +x_130 = l_Lean_Syntax_node3(x_84, x_79, x_2, x_77, x_3); +lean_inc(x_84); +x_131 = l_Lean_Syntax_node2(x_84, x_122, x_129, x_130); +lean_inc(x_84); +x_132 = l_Lean_Syntax_node2(x_84, x_118, x_120, x_131); +lean_inc(x_85); +lean_inc(x_84); +x_133 = l_Lean_Syntax_node2(x_84, x_116, x_85, x_132); x_134 = 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__2_spec__2___closed__18; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_135 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_134); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_135 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_134); x_136 = 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__2_spec__2___closed__19; -lean_inc(x_83); +lean_inc(x_84); x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_83); +lean_ctor_set(x_137, 0, x_84); lean_ctor_set(x_137, 1, 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__2_spec__2___closed__20; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_139 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_138); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_139 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_138); x_140 = 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__2_spec__2___closed__21; -lean_inc(x_83); +lean_inc(x_84); x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_83); +lean_ctor_set(x_141, 0, x_84); lean_ctor_set(x_141, 1, x_140); lean_inc(x_4); lean_inc(x_79); -lean_inc(x_83); -x_142 = l_Lean_Syntax_node1(x_83, x_79, x_4); +lean_inc(x_84); +x_142 = l_Lean_Syntax_node1(x_84, x_79, x_4); x_143 = 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__2_spec__2___closed__22; -lean_inc(x_83); +lean_inc(x_84); x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_83); +lean_ctor_set(x_144, 0, x_84); lean_ctor_set(x_144, 1, x_143); -lean_inc(x_83); -x_145 = l_Lean_Syntax_node3(x_83, x_139, x_141, x_142, x_144); +lean_inc(x_84); +x_145 = l_Lean_Syntax_node3(x_84, x_139, x_141, x_142, x_144); x_146 = 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__2_spec__2___closed__23; x_147 = 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__2_spec__2___closed__24; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_148 = l_Lean_Name_mkStr4(x_92, x_90, x_146, x_147); -lean_inc_n(x_80, 2); -lean_inc(x_83); -x_149 = l_Lean_Syntax_node2(x_83, x_148, x_80, x_80); -lean_inc(x_80); -lean_inc(x_83); -x_150 = l_Lean_Syntax_node4(x_83, x_135, x_137, x_145, x_149, x_80); -lean_inc(x_83); -x_151 = l_Lean_Syntax_node6(x_83, x_106, x_110, x_98, x_80, x_114, x_133, x_150); -x_152 = l_Lean_Syntax_node2(x_83, x_85, x_104, x_151); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_148 = l_Lean_Name_mkStr4(x_80, x_89, x_146, x_147); +lean_inc_n(x_85, 2); +lean_inc(x_84); +x_149 = l_Lean_Syntax_node2(x_84, x_148, x_85, x_85); +lean_inc(x_85); +lean_inc(x_84); +x_150 = l_Lean_Syntax_node4(x_84, x_135, x_137, x_145, x_149, x_85); +lean_inc(x_84); +x_151 = l_Lean_Syntax_node6(x_84, x_106, x_110, x_98, x_85, x_114, x_133, x_150); +x_152 = l_Lean_Syntax_node2(x_84, x_88, x_104, x_151); x_153 = lean_array_push(x_15, x_152); x_154 = 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__2_spec__2___closed__25; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_155 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_154); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_155 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_154); x_156 = 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__2_spec__2___closed__26; lean_inc(x_100); x_157 = lean_alloc_ctor(2, 2, 0); @@ -3577,9 +3577,9 @@ lean_ctor_set(x_157, 0, x_100); lean_ctor_set(x_157, 1, x_156); x_158 = 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__2_spec__2___closed__28; x_159 = 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__2_spec__2___closed__29; -lean_inc(x_89); -lean_inc(x_88); -x_160 = l_Lean_addMacroScope(x_88, x_159, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_160 = l_Lean_addMacroScope(x_94, x_159, x_78); lean_inc(x_100); x_161 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_161, 0, x_100); @@ -3587,9 +3587,9 @@ lean_ctor_set(x_161, 1, x_158); lean_ctor_set(x_161, 2, x_160); lean_ctor_set(x_161, 3, x_127); x_162 = 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__2_spec__2___closed__30; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_163 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_162); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_163 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_162); x_164 = 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__2_spec__2___closed__31; lean_inc(x_100); x_165 = lean_alloc_ctor(2, 2, 0); @@ -3602,22 +3602,22 @@ lean_ctor_set(x_166, 0, x_100); lean_ctor_set(x_166, 1, x_79); lean_ctor_set(x_166, 2, x_87); x_167 = 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__2_spec__2___closed__32; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_168 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_167); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_168 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_167); x_169 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_170 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_169); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_170 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_169); x_171 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_172 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_171); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_172 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_171); x_173 = 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__2_spec__2___closed__34; x_174 = 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__2_spec__2___closed__35; -lean_inc(x_89); -lean_inc(x_88); -x_175 = l_Lean_addMacroScope(x_88, x_174, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_175 = l_Lean_addMacroScope(x_94, x_174, x_78); lean_inc(x_100); x_176 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_176, 0, x_100); @@ -3629,9 +3629,9 @@ lean_inc(x_172); lean_inc(x_100); x_177 = l_Lean_Syntax_node2(x_100, x_172, x_176, x_166); x_178 = 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__2_spec__2___closed__36; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_179 = l_Lean_Name_mkStr4(x_92, x_90, x_107, x_178); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_179 = l_Lean_Name_mkStr4(x_80, x_89, x_107, x_178); lean_inc(x_100); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_100); @@ -3650,9 +3650,9 @@ lean_inc(x_100); x_183 = l_Lean_Syntax_node2(x_100, x_170, x_177, x_182); x_184 = 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__2_spec__2___closed__38; x_185 = 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__2_spec__2___closed__39; -lean_inc(x_89); -lean_inc(x_88); -x_186 = l_Lean_addMacroScope(x_88, x_185, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_186 = l_Lean_addMacroScope(x_94, x_185, x_78); lean_inc(x_100); x_187 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_187, 0, x_100); @@ -3678,9 +3678,9 @@ lean_inc(x_100); x_191 = l_Lean_Syntax_node2(x_100, x_170, x_188, x_190); x_192 = 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__2_spec__2___closed__41; x_193 = 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__2_spec__2___closed__42; -lean_inc(x_89); -lean_inc(x_88); -x_194 = l_Lean_addMacroScope(x_88, x_193, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_194 = l_Lean_addMacroScope(x_94, x_193, x_78); lean_inc(x_100); x_195 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_195, 0, x_100); @@ -3694,32 +3694,32 @@ if (x_1 == 0) { lean_object* x_197; x_197 = 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__2_spec__2___closed__46; -x_18 = x_79; -x_19 = x_166; -x_20 = x_180; -x_21 = x_127; -x_22 = x_165; -x_23 = x_179; -x_24 = x_157; -x_25 = x_168; -x_26 = x_88; -x_27 = x_191; -x_28 = x_183; -x_29 = x_89; -x_30 = x_90; -x_31 = x_92; -x_32 = x_170; -x_33 = x_100; -x_34 = x_196; -x_35 = x_119; -x_36 = x_153; -x_37 = x_161; -x_38 = x_101; +x_18 = x_100; +x_19 = x_155; +x_20 = x_157; +x_21 = x_183; +x_22 = x_163; +x_23 = x_168; +x_24 = x_89; +x_25 = x_179; +x_26 = x_153; +x_27 = x_127; +x_28 = x_119; +x_29 = x_161; +x_30 = x_78; +x_31 = x_79; +x_32 = x_80; +x_33 = x_166; +x_34 = x_191; +x_35 = x_196; +x_36 = x_127; +x_37 = x_170; +x_38 = x_165; x_39 = x_107; -x_40 = x_127; -x_41 = x_126; -x_42 = x_163; -x_43 = x_155; +x_40 = x_126; +x_41 = x_101; +x_42 = x_180; +x_43 = x_94; x_44 = x_197; goto block_74; } @@ -3727,32 +3727,32 @@ else { lean_object* x_198; x_198 = 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__2_spec__2___closed__49; -x_18 = x_79; -x_19 = x_166; -x_20 = x_180; -x_21 = x_127; -x_22 = x_165; -x_23 = x_179; -x_24 = x_157; -x_25 = x_168; -x_26 = x_88; -x_27 = x_191; -x_28 = x_183; -x_29 = x_89; -x_30 = x_90; -x_31 = x_92; -x_32 = x_170; -x_33 = x_100; -x_34 = x_196; -x_35 = x_119; -x_36 = x_153; -x_37 = x_161; -x_38 = x_101; +x_18 = x_100; +x_19 = x_155; +x_20 = x_157; +x_21 = x_183; +x_22 = x_163; +x_23 = x_168; +x_24 = x_89; +x_25 = x_179; +x_26 = x_153; +x_27 = x_127; +x_28 = x_119; +x_29 = x_161; +x_30 = x_78; +x_31 = x_79; +x_32 = x_80; +x_33 = x_166; +x_34 = x_191; +x_35 = x_196; +x_36 = x_127; +x_37 = x_170; +x_38 = x_165; x_39 = x_107; -x_40 = x_127; -x_41 = x_126; -x_42 = x_163; -x_43 = x_155; +x_40 = x_126; +x_41 = x_101; +x_42 = x_180; +x_43 = x_94; x_44 = x_198; goto block_74; } @@ -3769,71 +3769,71 @@ lean_inc_ref(x_87); x_201 = l_Array_append___redArg(x_87, x_95); lean_dec_ref(x_95); lean_inc(x_79); -lean_inc(x_83); +lean_inc(x_84); x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_83); +lean_ctor_set(x_202, 0, x_84); lean_ctor_set(x_202, 1, x_79); lean_ctor_set(x_202, 2, x_201); -lean_inc_n(x_80, 6); -lean_inc(x_83); -x_203 = l_Lean_Syntax_node7(x_83, x_94, x_80, x_80, x_202, x_80, x_80, x_80, x_80); +lean_inc_n(x_85, 6); +lean_inc(x_84); +x_203 = l_Lean_Syntax_node7(x_84, x_93, x_85, x_85, x_202, x_85, x_85, x_85, x_85); x_204 = 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__2_spec__2___closed__5; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_205 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_204); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_205 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_204); x_206 = l_Lake_configDecl___closed__26; x_207 = 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__2_spec__2___closed__6; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_208 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_207); -lean_inc(x_80); -lean_inc(x_83); -x_209 = l_Lean_Syntax_node1(x_83, x_208, x_80); -lean_inc(x_83); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_208 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_207); +lean_inc(x_85); +lean_inc(x_84); +x_209 = l_Lean_Syntax_node1(x_84, x_208, x_85); +lean_inc(x_84); x_210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_210, 0, x_83); +lean_ctor_set(x_210, 0, x_84); lean_ctor_set(x_210, 1, x_204); x_211 = l_Lake_configDecl___closed__8; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_212 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_211); -lean_inc(x_80); -lean_inc(x_83); -x_213 = l_Lean_Syntax_node2(x_83, x_212, x_81, x_80); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_212 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_211); +lean_inc(x_85); +lean_inc(x_84); +x_213 = l_Lean_Syntax_node2(x_84, x_212, x_86, x_85); lean_inc(x_79); -lean_inc(x_83); -x_214 = l_Lean_Syntax_node1(x_83, x_79, x_213); +lean_inc(x_84); +x_214 = l_Lean_Syntax_node1(x_84, x_79, x_213); x_215 = l_Lake_configField___closed__19; -lean_inc_ref(x_93); -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_216 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_215); +lean_inc_ref(x_91); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_216 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_215); x_217 = 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__2_spec__2___closed__7; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_218 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_217); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_218 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_217); x_219 = 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__2_spec__2___closed__8; -lean_inc(x_83); +lean_inc(x_84); x_220 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_220, 0, x_83); +lean_ctor_set(x_220, 0, x_84); lean_ctor_set(x_220, 1, x_219); x_221 = 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__2_spec__2___closed__9; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_222 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_221); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_222 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_221); x_223 = 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__2_spec__2___closed__11; x_224 = 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__2_spec__2___closed__12; -lean_inc(x_89); -lean_inc(x_88); -x_225 = l_Lean_addMacroScope(x_88, x_224, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_225 = l_Lean_addMacroScope(x_94, x_224, x_78); x_226 = l_Lake_configField___closed__1; x_227 = lean_box(0); x_228 = 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__2_spec__2___closed__17; -lean_inc(x_83); +lean_inc(x_84); x_229 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_229, 0, x_83); +lean_ctor_set(x_229, 0, x_84); lean_ctor_set(x_229, 1, x_223); lean_ctor_set(x_229, 2, x_225); lean_ctor_set(x_229, 3, x_228); @@ -3841,63 +3841,63 @@ lean_inc(x_3); lean_inc(x_77); lean_inc(x_2); lean_inc(x_79); -lean_inc(x_83); -x_230 = l_Lean_Syntax_node3(x_83, x_79, x_2, x_77, x_3); -lean_inc(x_83); -x_231 = l_Lean_Syntax_node2(x_83, x_222, x_229, x_230); -lean_inc(x_83); -x_232 = l_Lean_Syntax_node2(x_83, x_218, x_220, x_231); -lean_inc(x_80); -lean_inc(x_83); -x_233 = l_Lean_Syntax_node2(x_83, x_216, x_80, x_232); +lean_inc(x_84); +x_230 = l_Lean_Syntax_node3(x_84, x_79, x_2, x_77, x_3); +lean_inc(x_84); +x_231 = l_Lean_Syntax_node2(x_84, x_222, x_229, x_230); +lean_inc(x_84); +x_232 = l_Lean_Syntax_node2(x_84, x_218, x_220, x_231); +lean_inc(x_85); +lean_inc(x_84); +x_233 = l_Lean_Syntax_node2(x_84, x_216, x_85, x_232); x_234 = 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__2_spec__2___closed__18; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_235 = l_Lean_Name_mkStr4(x_92, x_90, x_93, x_234); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_235 = l_Lean_Name_mkStr4(x_80, x_89, x_91, x_234); x_236 = 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__2_spec__2___closed__19; -lean_inc(x_83); +lean_inc(x_84); x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_83); +lean_ctor_set(x_237, 0, x_84); lean_ctor_set(x_237, 1, x_236); x_238 = 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__2_spec__2___closed__20; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_239 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_238); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_239 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_238); x_240 = 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__2_spec__2___closed__21; -lean_inc(x_83); +lean_inc(x_84); x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_83); +lean_ctor_set(x_241, 0, x_84); lean_ctor_set(x_241, 1, x_240); lean_inc(x_4); lean_inc(x_79); -lean_inc(x_83); -x_242 = l_Lean_Syntax_node1(x_83, x_79, x_4); +lean_inc(x_84); +x_242 = l_Lean_Syntax_node1(x_84, x_79, x_4); x_243 = 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__2_spec__2___closed__22; -lean_inc(x_83); +lean_inc(x_84); x_244 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_244, 0, x_83); +lean_ctor_set(x_244, 0, x_84); lean_ctor_set(x_244, 1, x_243); -lean_inc(x_83); -x_245 = l_Lean_Syntax_node3(x_83, x_239, x_241, x_242, x_244); +lean_inc(x_84); +x_245 = l_Lean_Syntax_node3(x_84, x_239, x_241, x_242, x_244); x_246 = 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__2_spec__2___closed__23; x_247 = 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__2_spec__2___closed__24; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_248 = l_Lean_Name_mkStr4(x_92, x_90, x_246, x_247); -lean_inc_n(x_80, 2); -lean_inc(x_83); -x_249 = l_Lean_Syntax_node2(x_83, x_248, x_80, x_80); -lean_inc(x_80); -lean_inc(x_83); -x_250 = l_Lean_Syntax_node4(x_83, x_235, x_237, x_245, x_249, x_80); -lean_inc(x_83); -x_251 = l_Lean_Syntax_node6(x_83, x_205, x_209, x_210, x_80, x_214, x_233, x_250); -x_252 = l_Lean_Syntax_node2(x_83, x_85, x_203, x_251); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_248 = l_Lean_Name_mkStr4(x_80, x_89, x_246, x_247); +lean_inc_n(x_85, 2); +lean_inc(x_84); +x_249 = l_Lean_Syntax_node2(x_84, x_248, x_85, x_85); +lean_inc(x_85); +lean_inc(x_84); +x_250 = l_Lean_Syntax_node4(x_84, x_235, x_237, x_245, x_249, x_85); +lean_inc(x_84); +x_251 = l_Lean_Syntax_node6(x_84, x_205, x_209, x_210, x_85, x_214, x_233, x_250); +x_252 = l_Lean_Syntax_node2(x_84, x_88, x_203, x_251); x_253 = lean_array_push(x_15, x_252); x_254 = 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__2_spec__2___closed__25; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_255 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_254); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_255 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_254); x_256 = 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__2_spec__2___closed__26; lean_inc(x_199); x_257 = lean_alloc_ctor(2, 2, 0); @@ -3905,9 +3905,9 @@ lean_ctor_set(x_257, 0, x_199); lean_ctor_set(x_257, 1, x_256); x_258 = 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__2_spec__2___closed__28; x_259 = 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__2_spec__2___closed__29; -lean_inc(x_89); -lean_inc(x_88); -x_260 = l_Lean_addMacroScope(x_88, x_259, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_260 = l_Lean_addMacroScope(x_94, x_259, x_78); lean_inc(x_199); x_261 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_261, 0, x_199); @@ -3915,9 +3915,9 @@ lean_ctor_set(x_261, 1, x_258); lean_ctor_set(x_261, 2, x_260); lean_ctor_set(x_261, 3, x_227); x_262 = 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__2_spec__2___closed__30; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_263 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_262); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_263 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_262); x_264 = 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__2_spec__2___closed__31; lean_inc(x_199); x_265 = lean_alloc_ctor(2, 2, 0); @@ -3930,22 +3930,22 @@ lean_ctor_set(x_266, 0, x_199); lean_ctor_set(x_266, 1, x_79); lean_ctor_set(x_266, 2, x_87); x_267 = 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__2_spec__2___closed__32; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_268 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_267); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_268 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_267); x_269 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_270 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_269); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_270 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_269); x_271 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_272 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_271); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_272 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_271); x_273 = 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__2_spec__2___closed__34; x_274 = 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__2_spec__2___closed__35; -lean_inc(x_89); -lean_inc(x_88); -x_275 = l_Lean_addMacroScope(x_88, x_274, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_275 = l_Lean_addMacroScope(x_94, x_274, x_78); lean_inc(x_199); x_276 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_276, 0, x_199); @@ -3957,9 +3957,9 @@ lean_inc(x_272); lean_inc(x_199); x_277 = l_Lean_Syntax_node2(x_199, x_272, x_276, x_266); x_278 = 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__2_spec__2___closed__36; -lean_inc_ref(x_90); -lean_inc_ref(x_92); -x_279 = l_Lean_Name_mkStr4(x_92, x_90, x_206, x_278); +lean_inc_ref(x_89); +lean_inc_ref(x_80); +x_279 = l_Lean_Name_mkStr4(x_80, x_89, x_206, x_278); lean_inc(x_199); x_280 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_280, 0, x_199); @@ -3978,9 +3978,9 @@ lean_inc(x_199); x_283 = l_Lean_Syntax_node2(x_199, x_270, x_277, x_282); x_284 = 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__2_spec__2___closed__38; x_285 = 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__2_spec__2___closed__39; -lean_inc(x_89); -lean_inc(x_88); -x_286 = l_Lean_addMacroScope(x_88, x_285, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_286 = l_Lean_addMacroScope(x_94, x_285, x_78); lean_inc(x_199); x_287 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_287, 0, x_199); @@ -4006,9 +4006,9 @@ lean_inc(x_199); x_291 = l_Lean_Syntax_node2(x_199, x_270, x_288, x_290); x_292 = 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__2_spec__2___closed__41; x_293 = 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__2_spec__2___closed__42; -lean_inc(x_89); -lean_inc(x_88); -x_294 = l_Lean_addMacroScope(x_88, x_293, x_89); +lean_inc(x_78); +lean_inc(x_94); +x_294 = l_Lean_addMacroScope(x_94, x_293, x_78); lean_inc(x_199); x_295 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_295, 0, x_199); @@ -4022,32 +4022,32 @@ if (x_1 == 0) { lean_object* x_297; x_297 = 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__2_spec__2___closed__46; -x_18 = x_79; -x_19 = x_266; -x_20 = x_280; -x_21 = x_227; -x_22 = x_265; -x_23 = x_279; -x_24 = x_257; -x_25 = x_268; -x_26 = x_88; -x_27 = x_291; -x_28 = x_283; -x_29 = x_89; -x_30 = x_90; -x_31 = x_92; -x_32 = x_270; -x_33 = x_199; -x_34 = x_296; -x_35 = x_219; -x_36 = x_253; -x_37 = x_261; -x_38 = x_200; +x_18 = x_199; +x_19 = x_255; +x_20 = x_257; +x_21 = x_283; +x_22 = x_263; +x_23 = x_268; +x_24 = x_89; +x_25 = x_279; +x_26 = x_253; +x_27 = x_227; +x_28 = x_219; +x_29 = x_261; +x_30 = x_78; +x_31 = x_79; +x_32 = x_80; +x_33 = x_266; +x_34 = x_291; +x_35 = x_296; +x_36 = x_227; +x_37 = x_270; +x_38 = x_265; x_39 = x_206; -x_40 = x_227; -x_41 = x_226; -x_42 = x_263; -x_43 = x_255; +x_40 = x_226; +x_41 = x_200; +x_42 = x_280; +x_43 = x_94; x_44 = x_297; goto block_74; } @@ -4055,32 +4055,32 @@ else { lean_object* x_298; x_298 = 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__2_spec__2___closed__49; -x_18 = x_79; -x_19 = x_266; -x_20 = x_280; -x_21 = x_227; -x_22 = x_265; -x_23 = x_279; -x_24 = x_257; -x_25 = x_268; -x_26 = x_88; -x_27 = x_291; -x_28 = x_283; -x_29 = x_89; -x_30 = x_90; -x_31 = x_92; -x_32 = x_270; -x_33 = x_199; -x_34 = x_296; -x_35 = x_219; -x_36 = x_253; -x_37 = x_261; -x_38 = x_200; +x_18 = x_199; +x_19 = x_255; +x_20 = x_257; +x_21 = x_283; +x_22 = x_263; +x_23 = x_268; +x_24 = x_89; +x_25 = x_279; +x_26 = x_253; +x_27 = x_227; +x_28 = x_219; +x_29 = x_261; +x_30 = x_78; +x_31 = x_79; +x_32 = x_80; +x_33 = x_266; +x_34 = x_291; +x_35 = x_296; +x_36 = x_227; +x_37 = x_270; +x_38 = x_265; x_39 = x_206; -x_40 = x_227; -x_41 = x_226; -x_42 = x_263; -x_43 = x_255; +x_40 = x_226; +x_41 = x_200; +x_42 = x_280; +x_43 = x_94; x_44 = x_298; goto block_74; } @@ -4119,29 +4119,29 @@ if (lean_obj_tag(x_6) == 0) { lean_object* x_319; x_319 = 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__2_spec__2___closed__53; -lean_inc(x_301); -lean_inc(x_303); lean_inc(x_302); -lean_inc(x_306); -lean_inc(x_304); +lean_inc(x_301); lean_inc(x_305); -x_78 = x_305; +lean_inc(x_304); +lean_inc(x_306); +lean_inc(x_303); +x_78 = x_303; x_79 = x_316; -x_80 = x_318; -x_81 = x_310; -x_82 = x_304; -x_83 = x_308; -x_84 = x_309; -x_85 = x_314; -x_86 = x_306; +x_80 = x_311; +x_81 = x_309; +x_82 = x_306; +x_83 = x_304; +x_84 = x_308; +x_85 = x_318; +x_86 = x_310; x_87 = x_317; -x_88 = x_302; -x_89 = x_303; -x_90 = x_312; -x_91 = x_301; -x_92 = x_311; -x_93 = x_313; -x_94 = x_315; +x_88 = x_314; +x_89 = x_312; +x_90 = x_305; +x_91 = x_313; +x_92 = x_301; +x_93 = x_315; +x_94 = x_302; x_95 = x_319; goto block_299; } @@ -4151,29 +4151,29 @@ lean_object* x_320; lean_object* x_321; x_320 = lean_ctor_get(x_6, 0); lean_inc(x_320); x_321 = l_Array_mkArray1___redArg(x_320); -lean_inc(x_301); -lean_inc(x_303); lean_inc(x_302); -lean_inc(x_306); -lean_inc(x_304); +lean_inc(x_301); lean_inc(x_305); -x_78 = x_305; +lean_inc(x_304); +lean_inc(x_306); +lean_inc(x_303); +x_78 = x_303; x_79 = x_316; -x_80 = x_318; -x_81 = x_310; -x_82 = x_304; -x_83 = x_308; -x_84 = x_309; -x_85 = x_314; -x_86 = x_306; +x_80 = x_311; +x_81 = x_309; +x_82 = x_306; +x_83 = x_304; +x_84 = x_308; +x_85 = x_318; +x_86 = x_310; x_87 = x_317; -x_88 = x_302; -x_89 = x_303; -x_90 = x_312; -x_91 = x_301; -x_92 = x_311; -x_93 = x_313; -x_94 = x_315; +x_88 = x_314; +x_89 = x_312; +x_90 = x_305; +x_91 = x_313; +x_92 = x_301; +x_93 = x_315; +x_94 = x_302; x_95 = x_321; goto block_299; } @@ -4989,19 +4989,19 @@ goto block_1070; block_292: { lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = l_Lean_replaceRef(x_19, x_37); -lean_dec(x_37); +x_65 = l_Lean_replaceRef(x_19, x_34); +lean_dec(x_34); lean_inc(x_65); -lean_inc(x_31); -lean_inc(x_47); +lean_inc(x_59); +lean_inc(x_35); x_66 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_66, 0, x_55); -lean_ctor_set(x_66, 1, x_47); -lean_ctor_set(x_66, 2, x_31); -lean_ctor_set(x_66, 3, x_60); -lean_ctor_set(x_66, 4, x_59); +lean_ctor_set(x_66, 0, x_62); +lean_ctor_set(x_66, 1, x_35); +lean_ctor_set(x_66, 2, x_59); +lean_ctor_set(x_66, 3, x_43); +lean_ctor_set(x_66, 4, x_37); lean_ctor_set(x_66, 5, x_65); -x_67 = 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__2_spec__2___lam__1(x_17, x_65, x_66, x_58); +x_67 = 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__2_spec__2___lam__1(x_17, x_65, x_66, x_51); lean_dec_ref(x_66); lean_dec(x_65); x_68 = !lean_is_exclusive(x_67); @@ -5010,139 +5010,139 @@ if (x_68 == 0) lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; x_69 = lean_ctor_get(x_67, 0); x_70 = lean_ctor_get(x_67, 1); -lean_inc_ref(x_61); -x_71 = l_Array_append___redArg(x_61, x_64); +lean_inc_ref(x_53); +x_71 = l_Array_append___redArg(x_53, x_64); lean_dec_ref(x_64); -lean_inc(x_28); -lean_inc(x_50); +lean_inc(x_27); +lean_inc(x_63); x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_50); -lean_ctor_set(x_72, 1, x_28); +lean_ctor_set(x_72, 0, x_63); +lean_ctor_set(x_72, 1, x_27); lean_ctor_set(x_72, 2, x_71); -lean_inc_n(x_43, 6); -lean_inc(x_50); -x_73 = l_Lean_Syntax_node7(x_50, x_49, x_43, x_43, x_72, x_43, x_43, x_43, x_43); +lean_inc_n(x_50, 6); +lean_inc(x_63); +x_73 = l_Lean_Syntax_node7(x_63, x_56, x_50, x_50, x_72, x_50, x_50, x_50, x_50); x_74 = 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__2_spec__2___closed__5; -lean_inc_ref(x_52); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_75 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_74); +lean_inc_ref(x_60); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_75 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_74); x_76 = 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__2_spec__2___closed__6; +lean_inc_ref(x_48); lean_inc_ref(x_32); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_77 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_76); -lean_inc(x_43); -lean_inc(x_50); -x_78 = l_Lean_Syntax_node1(x_50, x_77, x_43); +lean_inc_ref(x_55); +x_77 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_76); lean_inc(x_50); +lean_inc(x_63); +x_78 = l_Lean_Syntax_node1(x_63, x_77, x_50); +lean_inc(x_63); lean_ctor_set_tag(x_67, 2); lean_ctor_set(x_67, 1, x_74); -lean_ctor_set(x_67, 0, x_50); -lean_inc(x_43); -lean_inc(x_50); -x_79 = l_Lean_Syntax_node2(x_50, x_34, x_56, x_43); -lean_inc(x_28); +lean_ctor_set(x_67, 0, x_63); lean_inc(x_50); -x_80 = l_Lean_Syntax_node1(x_50, x_28, x_79); +lean_inc(x_63); +x_79 = l_Lean_Syntax_node2(x_63, x_57, x_58, x_50); +lean_inc(x_27); +lean_inc(x_63); +x_80 = l_Lean_Syntax_node1(x_63, x_27, x_79); x_81 = l_Lake_configField___closed__19; +lean_inc_ref(x_60); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_82 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_81); lean_inc_ref(x_52); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_82 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_81); -lean_inc_ref(x_44); -lean_inc(x_50); +lean_inc(x_63); x_83 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_83, 0, x_50); -lean_ctor_set(x_83, 1, x_44); +lean_ctor_set(x_83, 0, x_63); +lean_ctor_set(x_83, 1, x_52); x_84 = 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__2_spec__2___closed__10; x_85 = 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__2_spec__2___closed__11; x_86 = 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__2_spec__2___closed__12; -lean_inc(x_31); -lean_inc(x_47); -x_87 = l_Lean_addMacroScope(x_47, x_86, x_31); -lean_inc_ref(x_46); -x_88 = l_Lean_Name_mkStr2(x_46, x_84); -lean_inc(x_63); +lean_inc(x_59); +lean_inc(x_35); +x_87 = l_Lean_addMacroScope(x_35, x_86, x_59); +lean_inc_ref(x_30); +x_88 = l_Lean_Name_mkStr2(x_30, x_84); +lean_inc(x_28); lean_inc(x_88); x_89 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_63); +lean_ctor_set(x_89, 1, x_28); x_90 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_90, 0, x_88); -lean_inc(x_42); +lean_inc(x_45); x_91 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_42); +lean_ctor_set(x_91, 1, x_45); x_92 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_92, 0, x_89); lean_ctor_set(x_92, 1, x_91); -lean_inc(x_50); +lean_inc(x_63); x_93 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_93, 0, x_50); +lean_ctor_set(x_93, 0, x_63); lean_ctor_set(x_93, 1, x_85); lean_ctor_set(x_93, 2, x_87); lean_ctor_set(x_93, 3, x_92); lean_inc(x_24); -lean_inc(x_51); +lean_inc(x_29); lean_inc(x_1); -lean_inc(x_28); -lean_inc(x_50); -x_94 = l_Lean_Syntax_node3(x_50, x_28, x_1, x_51, x_24); -lean_inc(x_50); -x_95 = l_Lean_Syntax_node2(x_50, x_53, x_93, x_94); -lean_inc(x_50); -x_96 = l_Lean_Syntax_node2(x_50, x_39, x_83, x_95); -lean_inc(x_43); +lean_inc(x_27); +lean_inc(x_63); +x_94 = l_Lean_Syntax_node3(x_63, x_27, x_1, x_29, x_24); +lean_inc(x_63); +x_95 = l_Lean_Syntax_node2(x_63, x_41, x_93, x_94); +lean_inc(x_63); +x_96 = l_Lean_Syntax_node2(x_63, x_42, x_83, x_95); lean_inc(x_50); -x_97 = l_Lean_Syntax_node2(x_50, x_82, x_43, x_96); +lean_inc(x_63); +x_97 = l_Lean_Syntax_node2(x_63, x_82, x_50, x_96); x_98 = 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__2_spec__2___closed__18; -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_99 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_98); -lean_inc_ref(x_62); -lean_inc(x_50); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_99 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_98); +lean_inc_ref(x_33); +lean_inc(x_63); x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_50); -lean_ctor_set(x_100, 1, x_62); +lean_ctor_set(x_100, 0, x_63); +lean_ctor_set(x_100, 1, x_33); x_101 = 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__2_spec__2___closed__20; +lean_inc_ref(x_48); lean_inc_ref(x_32); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_102 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_101); +lean_inc_ref(x_55); +x_102 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_101); x_103 = 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__2_spec__2___closed__21; -lean_inc(x_50); +lean_inc(x_63); x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_50); +lean_ctor_set(x_104, 0, x_63); lean_ctor_set(x_104, 1, x_103); -lean_inc(x_36); -lean_inc(x_28); -lean_inc(x_50); -x_105 = l_Lean_Syntax_node1(x_50, x_28, x_36); +lean_inc(x_61); +lean_inc(x_27); +lean_inc(x_63); +x_105 = l_Lean_Syntax_node1(x_63, x_27, x_61); x_106 = 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__2_spec__2___closed__22; -lean_inc(x_50); +lean_inc(x_63); x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_50); +lean_ctor_set(x_107, 0, x_63); lean_ctor_set(x_107, 1, x_106); -lean_inc(x_50); -x_108 = l_Lean_Syntax_node3(x_50, x_102, x_104, x_105, x_107); +lean_inc(x_63); +x_108 = l_Lean_Syntax_node3(x_63, x_102, x_104, x_105, x_107); x_109 = 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__2_spec__2___closed__23; x_110 = 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__2_spec__2___closed__24; -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_111 = l_Lean_Name_mkStr4(x_40, x_54, x_109, x_110); -lean_inc_n(x_43, 2); -lean_inc(x_50); -x_112 = l_Lean_Syntax_node2(x_50, x_111, x_43, x_43); -lean_inc(x_43); -lean_inc(x_50); -x_113 = l_Lean_Syntax_node4(x_50, x_99, x_100, x_108, x_112, x_43); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_111 = l_Lean_Name_mkStr4(x_55, x_32, x_109, x_110); +lean_inc_n(x_50, 2); +lean_inc(x_63); +x_112 = l_Lean_Syntax_node2(x_63, x_111, x_50, x_50); lean_inc(x_50); -x_114 = l_Lean_Syntax_node6(x_50, x_75, x_78, x_67, x_43, x_80, x_97, x_113); -x_115 = l_Lean_Syntax_node2(x_50, x_38, x_73, x_114); -x_116 = lean_array_push(x_27, x_115); +lean_inc(x_63); +x_113 = l_Lean_Syntax_node4(x_63, x_99, x_100, x_108, x_112, x_50); +lean_inc(x_63); +x_114 = l_Lean_Syntax_node6(x_63, x_75, x_78, x_67, x_50, x_80, x_97, x_113); +x_115 = l_Lean_Syntax_node2(x_63, x_47, x_73, x_114); +x_116 = lean_array_push(x_39, x_115); x_117 = 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__2_spec__2___closed__25; -x_118 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_117); +x_118 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_117); x_119 = 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__2_spec__2___closed__26; lean_inc(x_69); x_120 = lean_alloc_ctor(2, 2, 0); @@ -5150,131 +5150,131 @@ lean_ctor_set(x_120, 0, x_69); lean_ctor_set(x_120, 1, x_119); x_121 = 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__2_spec__2___closed__28; x_122 = 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__2_spec__2___closed__29; -lean_inc(x_31); -lean_inc(x_47); -x_123 = l_Lean_addMacroScope(x_47, x_122, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_123 = l_Lean_addMacroScope(x_35, x_122, x_59); +lean_inc(x_45); lean_inc(x_69); x_124 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_124, 0, x_69); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_42); +lean_ctor_set(x_124, 3, x_45); lean_inc(x_69); x_125 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_125, 0, x_69); -lean_ctor_set(x_125, 1, x_29); -lean_inc(x_28); +lean_ctor_set(x_125, 1, x_49); +lean_inc(x_27); lean_inc(x_69); x_126 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_126, 0, x_69); -lean_ctor_set(x_126, 1, x_28); -lean_ctor_set(x_126, 2, x_61); +lean_ctor_set(x_126, 1, x_27); +lean_ctor_set(x_126, 2, x_53); x_127 = 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__2_spec__2___closed__34; x_128 = 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__2_spec__2___closed__35; -lean_inc(x_31); -lean_inc(x_47); -x_129 = l_Lean_addMacroScope(x_47, x_128, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_129 = l_Lean_addMacroScope(x_35, x_128, x_59); +lean_inc(x_45); lean_inc(x_69); x_130 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_130, 0, x_69); lean_ctor_set(x_130, 1, x_127); lean_ctor_set(x_130, 2, x_129); -lean_ctor_set(x_130, 3, x_42); +lean_ctor_set(x_130, 3, x_45); lean_inc_ref(x_126); -lean_inc(x_41); +lean_inc(x_36); lean_inc(x_69); -x_131 = l_Lean_Syntax_node2(x_69, x_41, x_130, x_126); +x_131 = l_Lean_Syntax_node2(x_69, x_36, x_130, x_126); lean_inc(x_69); x_132 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_132, 0, x_69); -lean_ctor_set(x_132, 1, x_62); -lean_inc(x_51); +lean_ctor_set(x_132, 1, x_33); +lean_inc(x_29); lean_inc_ref(x_126); lean_inc_ref(x_132); -lean_inc(x_33); +lean_inc(x_46); lean_inc(x_69); -x_133 = l_Lean_Syntax_node3(x_69, x_33, x_132, x_126, x_51); +x_133 = l_Lean_Syntax_node3(x_69, x_46, x_132, x_126, x_29); lean_inc_ref_n(x_126, 2); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_69); -x_134 = l_Lean_Syntax_node3(x_69, x_28, x_126, x_126, x_133); +x_134 = l_Lean_Syntax_node3(x_69, x_27, x_126, x_126, x_133); lean_inc(x_134); -lean_inc(x_35); +lean_inc(x_38); lean_inc(x_69); -x_135 = l_Lean_Syntax_node2(x_69, x_35, x_131, x_134); +x_135 = l_Lean_Syntax_node2(x_69, x_38, x_131, x_134); x_136 = 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__2_spec__2___closed__38; x_137 = 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__2_spec__2___closed__39; -lean_inc(x_31); -lean_inc(x_47); -x_138 = l_Lean_addMacroScope(x_47, x_137, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_138 = l_Lean_addMacroScope(x_35, x_137, x_59); +lean_inc(x_45); lean_inc(x_69); x_139 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_139, 0, x_69); lean_ctor_set(x_139, 1, x_136); lean_ctor_set(x_139, 2, x_138); -lean_ctor_set(x_139, 3, x_42); +lean_ctor_set(x_139, 3, x_45); lean_inc_ref(x_126); -lean_inc(x_41); +lean_inc(x_36); lean_inc(x_69); -x_140 = l_Lean_Syntax_node2(x_69, x_41, x_139, x_126); -lean_inc(x_35); +x_140 = l_Lean_Syntax_node2(x_69, x_36, x_139, x_126); +lean_inc(x_38); lean_inc(x_69); -x_141 = l_Lean_Syntax_node2(x_69, x_35, x_140, x_134); +x_141 = l_Lean_Syntax_node2(x_69, x_38, x_140, x_134); x_142 = 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__2_spec__2___closed__41; x_143 = 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__2_spec__2___closed__42; -lean_inc(x_31); -lean_inc(x_47); -x_144 = l_Lean_addMacroScope(x_47, x_143, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_144 = l_Lean_addMacroScope(x_35, x_143, x_59); +lean_inc(x_45); lean_inc(x_69); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_69); lean_ctor_set(x_145, 1, x_142); lean_ctor_set(x_145, 2, x_144); -lean_ctor_set(x_145, 3, x_42); +lean_ctor_set(x_145, 3, x_45); lean_inc_ref(x_126); lean_inc(x_69); -x_146 = l_Lean_Syntax_node2(x_69, x_41, x_145, x_126); +x_146 = l_Lean_Syntax_node2(x_69, x_36, x_145, x_126); x_147 = 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__2_spec__2___closed__49; lean_inc_ref(x_126); lean_inc(x_69); -x_148 = l_Lean_Syntax_node3(x_69, x_33, x_132, x_126, x_147); +x_148 = l_Lean_Syntax_node3(x_69, x_46, x_132, x_126, x_147); lean_inc_ref_n(x_126, 2); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_69); -x_149 = l_Lean_Syntax_node3(x_69, x_28, x_126, x_126, x_148); +x_149 = l_Lean_Syntax_node3(x_69, x_27, x_126, x_126, x_148); lean_inc(x_69); -x_150 = l_Lean_Syntax_node2(x_69, x_35, x_146, x_149); +x_150 = l_Lean_Syntax_node2(x_69, x_38, x_146, x_149); lean_inc_ref_n(x_126, 3); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_69); -x_151 = l_Lean_Syntax_node6(x_69, x_28, x_135, x_126, x_141, x_126, x_150, x_126); +x_151 = l_Lean_Syntax_node6(x_69, x_27, x_135, x_126, x_141, x_126, x_150, x_126); lean_inc(x_69); -x_152 = l_Lean_Syntax_node1(x_69, x_30, x_151); +x_152 = l_Lean_Syntax_node1(x_69, x_54, x_151); lean_inc_ref(x_126); lean_inc(x_69); -x_153 = l_Lean_Syntax_node1(x_69, x_48, x_126); +x_153 = l_Lean_Syntax_node1(x_69, x_44, x_126); lean_inc(x_69); x_154 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_154, 0, x_69); -lean_ctor_set(x_154, 1, x_44); +lean_ctor_set(x_154, 1, x_52); x_155 = 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__2_spec__2___closed__1; x_156 = 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__2_spec__2___closed__2; x_157 = 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__2_spec__2___closed__3; -x_158 = l_Lean_addMacroScope(x_47, x_157, x_31); -x_159 = l_Lean_Name_mkStr2(x_46, x_155); +x_158 = l_Lean_addMacroScope(x_35, x_157, x_59); +x_159 = l_Lean_Name_mkStr2(x_30, x_155); lean_inc(x_159); x_160 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_63); +lean_ctor_set(x_160, 1, x_28); x_161 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_161, 0, x_159); x_162 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_42); +lean_ctor_set(x_162, 1, x_45); x_163 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_163, 0, x_160); lean_ctor_set(x_163, 1, x_162); @@ -5284,17 +5284,17 @@ lean_ctor_set(x_164, 0, x_69); lean_ctor_set(x_164, 1, x_156); lean_ctor_set(x_164, 2, x_158); lean_ctor_set(x_164, 3, x_163); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_69); -x_165 = l_Lean_Syntax_node2(x_69, x_28, x_154, x_164); +x_165 = l_Lean_Syntax_node2(x_69, x_27, x_154, x_164); lean_inc(x_69); x_166 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_166, 0, x_69); -lean_ctor_set(x_166, 1, x_45); +lean_ctor_set(x_166, 1, x_31); lean_inc(x_69); -x_167 = l_Lean_Syntax_node6(x_69, x_57, x_125, x_126, x_152, x_153, x_165, x_166); +x_167 = l_Lean_Syntax_node6(x_69, x_40, x_125, x_126, x_152, x_153, x_165, x_166); lean_inc(x_69); -x_168 = l_Lean_Syntax_node1(x_69, x_28, x_167); +x_168 = l_Lean_Syntax_node1(x_69, x_27, x_167); x_169 = l_Lean_Syntax_node4(x_69, x_118, x_19, x_120, x_124, x_168); if (lean_is_scalar(x_20)) { x_170 = lean_alloc_ctor(0, 2, 0); @@ -5309,8 +5309,8 @@ x_173 = lean_nat_dec_lt(x_171, x_172); if (x_173 == 0) { lean_dec(x_172); -lean_dec(x_51); -lean_dec(x_36); +lean_dec(x_61); +lean_dec(x_29); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_170; @@ -5324,8 +5324,8 @@ x_174 = lean_nat_dec_le(x_172, x_172); if (x_174 == 0) { lean_dec(x_172); -lean_dec(x_51); -lean_dec(x_36); +lean_dec(x_61); +lean_dec(x_29); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_170; @@ -5341,7 +5341,7 @@ lean_dec(x_172); lean_inc_ref(x_9); lean_inc(x_3); lean_inc(x_1); -x_177 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_36, x_51, x_3, x_4, x_23, x_175, x_176, x_170, x_9, x_70); +x_177 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_61, x_29, x_3, x_4, x_23, x_175, x_176, x_170, x_9, x_70); lean_dec_ref(x_23); x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); @@ -5362,139 +5362,139 @@ x_181 = lean_ctor_get(x_67, 1); lean_inc(x_181); lean_inc(x_180); lean_dec(x_67); -lean_inc_ref(x_61); -x_182 = l_Array_append___redArg(x_61, x_64); +lean_inc_ref(x_53); +x_182 = l_Array_append___redArg(x_53, x_64); lean_dec_ref(x_64); -lean_inc(x_28); -lean_inc(x_50); +lean_inc(x_27); +lean_inc(x_63); x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_50); -lean_ctor_set(x_183, 1, x_28); +lean_ctor_set(x_183, 0, x_63); +lean_ctor_set(x_183, 1, x_27); lean_ctor_set(x_183, 2, x_182); -lean_inc_n(x_43, 6); -lean_inc(x_50); -x_184 = l_Lean_Syntax_node7(x_50, x_49, x_43, x_43, x_183, x_43, x_43, x_43, x_43); +lean_inc_n(x_50, 6); +lean_inc(x_63); +x_184 = l_Lean_Syntax_node7(x_63, x_56, x_50, x_50, x_183, x_50, x_50, x_50, x_50); x_185 = 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__2_spec__2___closed__5; -lean_inc_ref(x_52); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_186 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_185); +lean_inc_ref(x_60); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_186 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_185); x_187 = 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__2_spec__2___closed__6; +lean_inc_ref(x_48); lean_inc_ref(x_32); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_188 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_187); -lean_inc(x_43); -lean_inc(x_50); -x_189 = l_Lean_Syntax_node1(x_50, x_188, x_43); +lean_inc_ref(x_55); +x_188 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_187); lean_inc(x_50); +lean_inc(x_63); +x_189 = l_Lean_Syntax_node1(x_63, x_188, x_50); +lean_inc(x_63); x_190 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_190, 0, x_50); +lean_ctor_set(x_190, 0, x_63); lean_ctor_set(x_190, 1, x_185); -lean_inc(x_43); -lean_inc(x_50); -x_191 = l_Lean_Syntax_node2(x_50, x_34, x_56, x_43); -lean_inc(x_28); lean_inc(x_50); -x_192 = l_Lean_Syntax_node1(x_50, x_28, x_191); +lean_inc(x_63); +x_191 = l_Lean_Syntax_node2(x_63, x_57, x_58, x_50); +lean_inc(x_27); +lean_inc(x_63); +x_192 = l_Lean_Syntax_node1(x_63, x_27, x_191); x_193 = l_Lake_configField___closed__19; +lean_inc_ref(x_60); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_194 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_193); lean_inc_ref(x_52); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_194 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_193); -lean_inc_ref(x_44); -lean_inc(x_50); +lean_inc(x_63); x_195 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_195, 0, x_50); -lean_ctor_set(x_195, 1, x_44); +lean_ctor_set(x_195, 0, x_63); +lean_ctor_set(x_195, 1, x_52); x_196 = 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__2_spec__2___closed__10; x_197 = 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__2_spec__2___closed__11; x_198 = 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__2_spec__2___closed__12; -lean_inc(x_31); -lean_inc(x_47); -x_199 = l_Lean_addMacroScope(x_47, x_198, x_31); -lean_inc_ref(x_46); -x_200 = l_Lean_Name_mkStr2(x_46, x_196); -lean_inc(x_63); +lean_inc(x_59); +lean_inc(x_35); +x_199 = l_Lean_addMacroScope(x_35, x_198, x_59); +lean_inc_ref(x_30); +x_200 = l_Lean_Name_mkStr2(x_30, x_196); +lean_inc(x_28); lean_inc(x_200); x_201 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_63); +lean_ctor_set(x_201, 1, x_28); x_202 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_202, 0, x_200); -lean_inc(x_42); +lean_inc(x_45); x_203 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_42); +lean_ctor_set(x_203, 1, x_45); x_204 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_204, 0, x_201); lean_ctor_set(x_204, 1, x_203); -lean_inc(x_50); +lean_inc(x_63); x_205 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_205, 0, x_50); +lean_ctor_set(x_205, 0, x_63); lean_ctor_set(x_205, 1, x_197); lean_ctor_set(x_205, 2, x_199); lean_ctor_set(x_205, 3, x_204); lean_inc(x_24); -lean_inc(x_51); +lean_inc(x_29); lean_inc(x_1); -lean_inc(x_28); -lean_inc(x_50); -x_206 = l_Lean_Syntax_node3(x_50, x_28, x_1, x_51, x_24); -lean_inc(x_50); -x_207 = l_Lean_Syntax_node2(x_50, x_53, x_205, x_206); -lean_inc(x_50); -x_208 = l_Lean_Syntax_node2(x_50, x_39, x_195, x_207); -lean_inc(x_43); +lean_inc(x_27); +lean_inc(x_63); +x_206 = l_Lean_Syntax_node3(x_63, x_27, x_1, x_29, x_24); +lean_inc(x_63); +x_207 = l_Lean_Syntax_node2(x_63, x_41, x_205, x_206); +lean_inc(x_63); +x_208 = l_Lean_Syntax_node2(x_63, x_42, x_195, x_207); lean_inc(x_50); -x_209 = l_Lean_Syntax_node2(x_50, x_194, x_43, x_208); +lean_inc(x_63); +x_209 = l_Lean_Syntax_node2(x_63, x_194, x_50, x_208); x_210 = 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__2_spec__2___closed__18; -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_211 = l_Lean_Name_mkStr4(x_40, x_54, x_52, x_210); -lean_inc_ref(x_62); -lean_inc(x_50); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_211 = l_Lean_Name_mkStr4(x_55, x_32, x_60, x_210); +lean_inc_ref(x_33); +lean_inc(x_63); x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_50); -lean_ctor_set(x_212, 1, x_62); +lean_ctor_set(x_212, 0, x_63); +lean_ctor_set(x_212, 1, x_33); x_213 = 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__2_spec__2___closed__20; +lean_inc_ref(x_48); lean_inc_ref(x_32); -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_214 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_213); +lean_inc_ref(x_55); +x_214 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_213); x_215 = 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__2_spec__2___closed__21; -lean_inc(x_50); +lean_inc(x_63); x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_50); +lean_ctor_set(x_216, 0, x_63); lean_ctor_set(x_216, 1, x_215); -lean_inc(x_36); -lean_inc(x_28); -lean_inc(x_50); -x_217 = l_Lean_Syntax_node1(x_50, x_28, x_36); +lean_inc(x_61); +lean_inc(x_27); +lean_inc(x_63); +x_217 = l_Lean_Syntax_node1(x_63, x_27, x_61); x_218 = 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__2_spec__2___closed__22; -lean_inc(x_50); +lean_inc(x_63); x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_50); +lean_ctor_set(x_219, 0, x_63); lean_ctor_set(x_219, 1, x_218); -lean_inc(x_50); -x_220 = l_Lean_Syntax_node3(x_50, x_214, x_216, x_217, x_219); +lean_inc(x_63); +x_220 = l_Lean_Syntax_node3(x_63, x_214, x_216, x_217, x_219); x_221 = 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__2_spec__2___closed__23; x_222 = 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__2_spec__2___closed__24; -lean_inc_ref(x_54); -lean_inc_ref(x_40); -x_223 = l_Lean_Name_mkStr4(x_40, x_54, x_221, x_222); -lean_inc_n(x_43, 2); -lean_inc(x_50); -x_224 = l_Lean_Syntax_node2(x_50, x_223, x_43, x_43); -lean_inc(x_43); -lean_inc(x_50); -x_225 = l_Lean_Syntax_node4(x_50, x_211, x_212, x_220, x_224, x_43); +lean_inc_ref(x_32); +lean_inc_ref(x_55); +x_223 = l_Lean_Name_mkStr4(x_55, x_32, x_221, x_222); +lean_inc_n(x_50, 2); +lean_inc(x_63); +x_224 = l_Lean_Syntax_node2(x_63, x_223, x_50, x_50); lean_inc(x_50); -x_226 = l_Lean_Syntax_node6(x_50, x_186, x_189, x_190, x_43, x_192, x_209, x_225); -x_227 = l_Lean_Syntax_node2(x_50, x_38, x_184, x_226); -x_228 = lean_array_push(x_27, x_227); +lean_inc(x_63); +x_225 = l_Lean_Syntax_node4(x_63, x_211, x_212, x_220, x_224, x_50); +lean_inc(x_63); +x_226 = l_Lean_Syntax_node6(x_63, x_186, x_189, x_190, x_50, x_192, x_209, x_225); +x_227 = l_Lean_Syntax_node2(x_63, x_47, x_184, x_226); +x_228 = lean_array_push(x_39, x_227); x_229 = 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__2_spec__2___closed__25; -x_230 = l_Lean_Name_mkStr4(x_40, x_54, x_32, x_229); +x_230 = l_Lean_Name_mkStr4(x_55, x_32, x_48, x_229); x_231 = 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__2_spec__2___closed__26; lean_inc(x_180); x_232 = lean_alloc_ctor(2, 2, 0); @@ -5502,131 +5502,131 @@ lean_ctor_set(x_232, 0, x_180); lean_ctor_set(x_232, 1, x_231); x_233 = 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__2_spec__2___closed__28; x_234 = 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__2_spec__2___closed__29; -lean_inc(x_31); -lean_inc(x_47); -x_235 = l_Lean_addMacroScope(x_47, x_234, x_31); -lean_inc(x_42); -lean_inc(x_180); +lean_inc(x_59); +lean_inc(x_35); +x_235 = l_Lean_addMacroScope(x_35, x_234, x_59); +lean_inc(x_45); +lean_inc(x_180); x_236 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_236, 0, x_180); lean_ctor_set(x_236, 1, x_233); lean_ctor_set(x_236, 2, x_235); -lean_ctor_set(x_236, 3, x_42); +lean_ctor_set(x_236, 3, x_45); lean_inc(x_180); x_237 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_237, 0, x_180); -lean_ctor_set(x_237, 1, x_29); -lean_inc(x_28); +lean_ctor_set(x_237, 1, x_49); +lean_inc(x_27); lean_inc(x_180); x_238 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_238, 0, x_180); -lean_ctor_set(x_238, 1, x_28); -lean_ctor_set(x_238, 2, x_61); +lean_ctor_set(x_238, 1, x_27); +lean_ctor_set(x_238, 2, x_53); x_239 = 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__2_spec__2___closed__34; x_240 = 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__2_spec__2___closed__35; -lean_inc(x_31); -lean_inc(x_47); -x_241 = l_Lean_addMacroScope(x_47, x_240, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_241 = l_Lean_addMacroScope(x_35, x_240, x_59); +lean_inc(x_45); lean_inc(x_180); x_242 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_242, 0, x_180); lean_ctor_set(x_242, 1, x_239); lean_ctor_set(x_242, 2, x_241); -lean_ctor_set(x_242, 3, x_42); +lean_ctor_set(x_242, 3, x_45); lean_inc_ref(x_238); -lean_inc(x_41); +lean_inc(x_36); lean_inc(x_180); -x_243 = l_Lean_Syntax_node2(x_180, x_41, x_242, x_238); +x_243 = l_Lean_Syntax_node2(x_180, x_36, x_242, x_238); lean_inc(x_180); x_244 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_244, 0, x_180); -lean_ctor_set(x_244, 1, x_62); -lean_inc(x_51); +lean_ctor_set(x_244, 1, x_33); +lean_inc(x_29); lean_inc_ref(x_238); lean_inc_ref(x_244); -lean_inc(x_33); +lean_inc(x_46); lean_inc(x_180); -x_245 = l_Lean_Syntax_node3(x_180, x_33, x_244, x_238, x_51); +x_245 = l_Lean_Syntax_node3(x_180, x_46, x_244, x_238, x_29); lean_inc_ref_n(x_238, 2); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_180); -x_246 = l_Lean_Syntax_node3(x_180, x_28, x_238, x_238, x_245); +x_246 = l_Lean_Syntax_node3(x_180, x_27, x_238, x_238, x_245); lean_inc(x_246); -lean_inc(x_35); +lean_inc(x_38); lean_inc(x_180); -x_247 = l_Lean_Syntax_node2(x_180, x_35, x_243, x_246); +x_247 = l_Lean_Syntax_node2(x_180, x_38, x_243, x_246); x_248 = 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__2_spec__2___closed__38; x_249 = 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__2_spec__2___closed__39; -lean_inc(x_31); -lean_inc(x_47); -x_250 = l_Lean_addMacroScope(x_47, x_249, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_250 = l_Lean_addMacroScope(x_35, x_249, x_59); +lean_inc(x_45); lean_inc(x_180); x_251 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_251, 0, x_180); lean_ctor_set(x_251, 1, x_248); lean_ctor_set(x_251, 2, x_250); -lean_ctor_set(x_251, 3, x_42); +lean_ctor_set(x_251, 3, x_45); lean_inc_ref(x_238); -lean_inc(x_41); +lean_inc(x_36); lean_inc(x_180); -x_252 = l_Lean_Syntax_node2(x_180, x_41, x_251, x_238); -lean_inc(x_35); +x_252 = l_Lean_Syntax_node2(x_180, x_36, x_251, x_238); +lean_inc(x_38); lean_inc(x_180); -x_253 = l_Lean_Syntax_node2(x_180, x_35, x_252, x_246); +x_253 = l_Lean_Syntax_node2(x_180, x_38, x_252, x_246); x_254 = 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__2_spec__2___closed__41; x_255 = 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__2_spec__2___closed__42; -lean_inc(x_31); -lean_inc(x_47); -x_256 = l_Lean_addMacroScope(x_47, x_255, x_31); -lean_inc(x_42); +lean_inc(x_59); +lean_inc(x_35); +x_256 = l_Lean_addMacroScope(x_35, x_255, x_59); +lean_inc(x_45); lean_inc(x_180); x_257 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_257, 0, x_180); lean_ctor_set(x_257, 1, x_254); lean_ctor_set(x_257, 2, x_256); -lean_ctor_set(x_257, 3, x_42); +lean_ctor_set(x_257, 3, x_45); lean_inc_ref(x_238); lean_inc(x_180); -x_258 = l_Lean_Syntax_node2(x_180, x_41, x_257, x_238); +x_258 = l_Lean_Syntax_node2(x_180, x_36, x_257, x_238); x_259 = 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__2_spec__2___closed__49; lean_inc_ref(x_238); lean_inc(x_180); -x_260 = l_Lean_Syntax_node3(x_180, x_33, x_244, x_238, x_259); +x_260 = l_Lean_Syntax_node3(x_180, x_46, x_244, x_238, x_259); lean_inc_ref_n(x_238, 2); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_180); -x_261 = l_Lean_Syntax_node3(x_180, x_28, x_238, x_238, x_260); +x_261 = l_Lean_Syntax_node3(x_180, x_27, x_238, x_238, x_260); lean_inc(x_180); -x_262 = l_Lean_Syntax_node2(x_180, x_35, x_258, x_261); +x_262 = l_Lean_Syntax_node2(x_180, x_38, x_258, x_261); lean_inc_ref_n(x_238, 3); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_180); -x_263 = l_Lean_Syntax_node6(x_180, x_28, x_247, x_238, x_253, x_238, x_262, x_238); +x_263 = l_Lean_Syntax_node6(x_180, x_27, x_247, x_238, x_253, x_238, x_262, x_238); lean_inc(x_180); -x_264 = l_Lean_Syntax_node1(x_180, x_30, x_263); +x_264 = l_Lean_Syntax_node1(x_180, x_54, x_263); lean_inc_ref(x_238); lean_inc(x_180); -x_265 = l_Lean_Syntax_node1(x_180, x_48, x_238); +x_265 = l_Lean_Syntax_node1(x_180, x_44, x_238); lean_inc(x_180); x_266 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_266, 0, x_180); -lean_ctor_set(x_266, 1, x_44); +lean_ctor_set(x_266, 1, x_52); x_267 = 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__2_spec__2___closed__1; x_268 = 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__2_spec__2___closed__2; x_269 = 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__2_spec__2___closed__3; -x_270 = l_Lean_addMacroScope(x_47, x_269, x_31); -x_271 = l_Lean_Name_mkStr2(x_46, x_267); +x_270 = l_Lean_addMacroScope(x_35, x_269, x_59); +x_271 = l_Lean_Name_mkStr2(x_30, x_267); lean_inc(x_271); x_272 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_63); +lean_ctor_set(x_272, 1, x_28); x_273 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_273, 0, x_271); x_274 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_42); +lean_ctor_set(x_274, 1, x_45); x_275 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_275, 0, x_272); lean_ctor_set(x_275, 1, x_274); @@ -5636,17 +5636,17 @@ lean_ctor_set(x_276, 0, x_180); lean_ctor_set(x_276, 1, x_268); lean_ctor_set(x_276, 2, x_270); lean_ctor_set(x_276, 3, x_275); -lean_inc(x_28); +lean_inc(x_27); lean_inc(x_180); -x_277 = l_Lean_Syntax_node2(x_180, x_28, x_266, x_276); +x_277 = l_Lean_Syntax_node2(x_180, x_27, x_266, x_276); lean_inc(x_180); x_278 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_278, 0, x_180); -lean_ctor_set(x_278, 1, x_45); +lean_ctor_set(x_278, 1, x_31); lean_inc(x_180); -x_279 = l_Lean_Syntax_node6(x_180, x_57, x_237, x_238, x_264, x_265, x_277, x_278); +x_279 = l_Lean_Syntax_node6(x_180, x_40, x_237, x_238, x_264, x_265, x_277, x_278); lean_inc(x_180); -x_280 = l_Lean_Syntax_node1(x_180, x_28, x_279); +x_280 = l_Lean_Syntax_node1(x_180, x_27, x_279); x_281 = l_Lean_Syntax_node4(x_180, x_230, x_19, x_232, x_236, x_280); if (lean_is_scalar(x_20)) { x_282 = lean_alloc_ctor(0, 2, 0); @@ -5661,8 +5661,8 @@ x_285 = lean_nat_dec_lt(x_283, x_284); if (x_285 == 0) { lean_dec(x_284); -lean_dec(x_51); -lean_dec(x_36); +lean_dec(x_61); +lean_dec(x_29); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_282; @@ -5676,8 +5676,8 @@ x_286 = lean_nat_dec_le(x_284, x_284); if (x_286 == 0) { lean_dec(x_284); -lean_dec(x_51); -lean_dec(x_36); +lean_dec(x_61); +lean_dec(x_29); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_282; @@ -5693,7 +5693,7 @@ lean_dec(x_284); lean_inc_ref(x_9); lean_inc(x_3); lean_inc(x_1); -x_289 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_36, x_51, x_3, x_4, x_23, x_287, x_288, x_282, x_9, x_181); +x_289 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_61, x_29, x_3, x_4, x_23, x_287, x_288, x_282, x_9, x_181); lean_dec_ref(x_23); x_290 = lean_ctor_get(x_289, 0); lean_inc(x_290); @@ -5710,7 +5710,7 @@ goto block_16; block_335: { lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; -x_327 = 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__2_spec__2___lam__1(x_17, x_303, x_9, x_10); +x_327 = 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__2_spec__2___lam__1(x_17, x_300, x_9, x_10); x_328 = lean_ctor_get(x_327, 0); lean_inc(x_328); x_329 = lean_ctor_get(x_327, 1); @@ -5718,13 +5718,13 @@ lean_inc(x_329); lean_dec_ref(x_327); x_330 = l_Lean_mkIdentFrom(x_22, x_326, x_17); lean_dec(x_22); -lean_inc_ref(x_323); -lean_inc(x_294); +lean_inc_ref(x_317); +lean_inc(x_293); lean_inc(x_328); x_331 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_331, 0, x_328); -lean_ctor_set(x_331, 1, x_294); -lean_ctor_set(x_331, 2, x_323); +lean_ctor_set(x_331, 1, x_293); +lean_ctor_set(x_331, 2, x_317); if (lean_obj_tag(x_3) == 0) { lean_object* x_332; @@ -5732,40 +5732,40 @@ x_332 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__ x_27 = x_293; x_28 = x_294; x_29 = x_295; -x_30 = x_298; +x_30 = x_296; x_31 = x_297; -x_32 = x_296; +x_32 = x_298; x_33 = x_299; x_34 = x_300; -x_35 = x_302; -x_36 = x_301; +x_35 = x_301; +x_36 = x_302; x_37 = x_303; x_38 = x_304; x_39 = x_305; x_40 = x_306; x_41 = x_307; x_42 = x_308; -x_43 = x_331; -x_44 = x_309; -x_45 = x_310; -x_46 = x_311; -x_47 = x_312; -x_48 = x_313; -x_49 = x_314; -x_50 = x_328; -x_51 = x_315; -x_52 = x_317; -x_53 = x_316; -x_54 = x_318; +x_43 = x_309; +x_44 = x_310; +x_45 = x_311; +x_46 = x_312; +x_47 = x_313; +x_48 = x_314; +x_49 = x_315; +x_50 = x_331; +x_51 = x_329; +x_52 = x_316; +x_53 = x_317; +x_54 = x_320; x_55 = x_319; -x_56 = x_330; -x_57 = x_320; -x_58 = x_329; -x_59 = x_321; +x_56 = x_318; +x_57 = x_321; +x_58 = x_330; +x_59 = x_323; x_60 = x_322; -x_61 = x_323; -x_62 = x_325; -x_63 = x_324; +x_61 = x_325; +x_62 = x_324; +x_63 = x_328; x_64 = x_332; goto block_292; } @@ -5778,40 +5778,40 @@ x_334 = l_Array_mkArray1___redArg(x_333); x_27 = x_293; x_28 = x_294; x_29 = x_295; -x_30 = x_298; +x_30 = x_296; x_31 = x_297; -x_32 = x_296; +x_32 = x_298; x_33 = x_299; x_34 = x_300; -x_35 = x_302; -x_36 = x_301; +x_35 = x_301; +x_36 = x_302; x_37 = x_303; x_38 = x_304; x_39 = x_305; x_40 = x_306; x_41 = x_307; x_42 = x_308; -x_43 = x_331; -x_44 = x_309; -x_45 = x_310; -x_46 = x_311; -x_47 = x_312; -x_48 = x_313; -x_49 = x_314; -x_50 = x_328; -x_51 = x_315; -x_52 = x_317; -x_53 = x_316; -x_54 = x_318; +x_43 = x_309; +x_44 = x_310; +x_45 = x_311; +x_46 = x_312; +x_47 = x_313; +x_48 = x_314; +x_49 = x_315; +x_50 = x_331; +x_51 = x_329; +x_52 = x_316; +x_53 = x_317; +x_54 = x_320; x_55 = x_319; -x_56 = x_330; -x_57 = x_320; -x_58 = x_329; -x_59 = x_321; +x_56 = x_318; +x_57 = x_321; +x_58 = x_330; +x_59 = x_323; x_60 = x_322; -x_61 = x_323; -x_62 = x_325; -x_63 = x_324; +x_61 = x_325; +x_62 = x_324; +x_63 = x_328; x_64 = x_334; goto block_292; } @@ -5819,21 +5819,21 @@ goto block_292; block_812: { lean_object* x_374; lean_object* x_375; lean_object* x_376; uint8_t x_377; -x_374 = l_Lean_replaceRef(x_19, x_347); +x_374 = l_Lean_replaceRef(x_19, x_344); lean_inc(x_374); -lean_inc(x_368); +lean_inc(x_348); +lean_inc(x_354); lean_inc(x_369); -lean_inc(x_341); -lean_inc(x_356); -lean_inc(x_365); +lean_inc(x_346); +lean_inc(x_372); x_375 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_375, 0, x_365); -lean_ctor_set(x_375, 1, x_356); -lean_ctor_set(x_375, 2, x_341); -lean_ctor_set(x_375, 3, x_369); -lean_ctor_set(x_375, 4, x_368); +lean_ctor_set(x_375, 0, x_372); +lean_ctor_set(x_375, 1, x_346); +lean_ctor_set(x_375, 2, x_369); +lean_ctor_set(x_375, 3, x_354); +lean_ctor_set(x_375, 4, x_348); lean_ctor_set(x_375, 5, x_374); -x_376 = 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__2_spec__2___lam__1(x_17, x_374, x_375, x_359); +x_376 = 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__2_spec__2___lam__1(x_17, x_374, x_375, x_368); lean_dec_ref(x_375); lean_dec(x_374); x_377 = !lean_is_exclusive(x_376); @@ -5843,36 +5843,36 @@ lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; x_378 = lean_ctor_get(x_376, 0); x_379 = lean_ctor_get(x_376, 1); x_380 = 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__2_spec__2___closed__25; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_381 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_380); +x_381 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_380); x_382 = 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__2_spec__2___closed__26; lean_inc(x_378); lean_ctor_set_tag(x_376, 2); lean_ctor_set(x_376, 1, x_382); x_383 = 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__4_spec__4___closed__1; x_384 = 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__4_spec__4___closed__2; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_385 = l_Lean_addMacroScope(x_346, x_384, x_369); lean_inc(x_356); -x_385 = l_Lean_addMacroScope(x_356, x_384, x_341); -lean_inc(x_352); lean_inc(x_378); x_386 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_386, 0, x_378); lean_ctor_set(x_386, 1, x_383); lean_ctor_set(x_386, 2, x_385); -lean_ctor_set(x_386, 3, x_352); +lean_ctor_set(x_386, 3, x_356); x_387 = 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__4_spec__4___closed__3; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_388 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_387); +x_388 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_387); x_389 = 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__4_spec__4___closed__4; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_390 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_389); +x_390 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_389); x_391 = 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__4_spec__4___closed__5; lean_inc(x_378); x_392 = lean_alloc_ctor(2, 2, 0); @@ -5881,37 +5881,37 @@ lean_ctor_set(x_392, 1, x_391); x_393 = 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__4_spec__4___closed__7; x_394 = 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__4_spec__4___closed__9; x_395 = lean_box(0); -lean_inc(x_341); -lean_inc(x_356); -x_396 = l_Lean_addMacroScope(x_356, x_395, x_341); -lean_inc_ref(x_355); -x_397 = l_Lean_Name_mkStr1(x_355); +lean_inc(x_369); +lean_inc(x_346); +x_396 = l_Lean_addMacroScope(x_346, x_395, x_369); +lean_inc_ref(x_341); +x_397 = l_Lean_Name_mkStr1(x_341); x_398 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_398, 0, x_397); -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_399 = l_Lean_Name_mkStr3(x_350, x_364, x_362); +x_399 = l_Lean_Name_mkStr3(x_364, x_343, x_367); x_400 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_400, 0, x_399); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_401 = l_Lean_Name_mkStr2(x_350, x_364); +x_401 = l_Lean_Name_mkStr2(x_364, x_343); x_402 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_402, 0, x_401); x_403 = 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__4_spec__4___closed__10; -lean_inc_ref(x_350); -x_404 = l_Lean_Name_mkStr2(x_350, x_403); +lean_inc_ref(x_364); +x_404 = l_Lean_Name_mkStr2(x_364, x_403); x_405 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_405, 0, x_404); -lean_inc_ref(x_350); -x_406 = l_Lean_Name_mkStr1(x_350); +lean_inc_ref(x_364); +x_406 = l_Lean_Name_mkStr1(x_364); x_407 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_407, 0, x_406); -lean_inc(x_352); +lean_inc(x_356); x_408 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_408, 0, x_407); -lean_ctor_set(x_408, 1, x_352); +lean_ctor_set(x_408, 1, x_356); x_409 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_409, 0, x_405); lean_ctor_set(x_409, 1, x_408); @@ -5938,19 +5938,19 @@ x_416 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__ x_417 = 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__4_spec__4___closed__13; x_418 = 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__4_spec__4___closed__14; x_419 = 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__4_spec__4___closed__15; -lean_inc(x_341); -lean_inc(x_356); -x_420 = l_Lean_addMacroScope(x_356, x_419, x_341); -lean_inc_ref(x_355); -x_421 = l_Lean_Name_mkStr3(x_355, x_417, x_418); -lean_inc(x_371); +lean_inc(x_369); +lean_inc(x_346); +x_420 = l_Lean_addMacroScope(x_346, x_419, x_369); +lean_inc_ref(x_341); +x_421 = l_Lean_Name_mkStr3(x_341, x_417, x_418); +lean_inc(x_339); x_422 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_422, 0, x_421); -lean_ctor_set(x_422, 1, x_371); -lean_inc(x_352); +lean_ctor_set(x_422, 1, x_339); +lean_inc(x_356); x_423 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_423, 0, x_422); -lean_ctor_set(x_423, 1, x_352); +lean_ctor_set(x_423, 1, x_356); lean_inc(x_378); x_424 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_424, 0, x_378); @@ -5958,12 +5958,12 @@ lean_ctor_set(x_424, 1, x_416); lean_ctor_set(x_424, 2, x_420); lean_ctor_set(x_424, 3, x_423); lean_inc(x_24); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_378); -x_425 = l_Lean_Syntax_node1(x_378, x_337, x_24); -lean_inc(x_363); +x_425 = l_Lean_Syntax_node1(x_378, x_338, x_24); +lean_inc(x_352); lean_inc(x_378); -x_426 = l_Lean_Syntax_node2(x_378, x_363, x_424, x_425); +x_426 = l_Lean_Syntax_node2(x_378, x_352, x_424, x_425); x_427 = 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__4_spec__4___closed__16; lean_inc(x_378); x_428 = lean_alloc_ctor(2, 2, 0); @@ -5971,22 +5971,22 @@ lean_ctor_set(x_428, 0, x_378); lean_ctor_set(x_428, 1, x_427); lean_inc(x_378); x_429 = l_Lean_Syntax_node3(x_378, x_388, x_415, x_426, x_428); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_378); -x_430 = l_Lean_Syntax_node1(x_378, x_337, x_429); +x_430 = l_Lean_Syntax_node1(x_378, x_338, x_429); lean_inc(x_381); x_431 = l_Lean_Syntax_node4(x_378, x_381, x_19, x_376, x_386, x_430); -x_432 = l_Lean_replaceRef(x_431, x_347); -lean_dec(x_347); +x_432 = l_Lean_replaceRef(x_431, x_344); +lean_dec(x_344); lean_inc(x_432); -lean_inc(x_341); -lean_inc(x_356); +lean_inc(x_369); +lean_inc(x_346); x_433 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_433, 0, x_365); -lean_ctor_set(x_433, 1, x_356); -lean_ctor_set(x_433, 2, x_341); -lean_ctor_set(x_433, 3, x_369); -lean_ctor_set(x_433, 4, x_368); +lean_ctor_set(x_433, 0, x_372); +lean_ctor_set(x_433, 1, x_346); +lean_ctor_set(x_433, 2, x_369); +lean_ctor_set(x_433, 3, x_354); +lean_ctor_set(x_433, 4, x_348); lean_ctor_set(x_433, 5, x_432); x_434 = 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__2_spec__2___lam__1(x_17, x_432, x_433, x_379); lean_dec_ref(x_433); @@ -5997,238 +5997,238 @@ if (x_435 == 0) lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; 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_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; x_436 = lean_ctor_get(x_434, 0); x_437 = lean_ctor_get(x_434, 1); -lean_inc_ref(x_370); -x_438 = l_Array_append___redArg(x_370, x_373); +lean_inc_ref(x_362); +x_438 = l_Array_append___redArg(x_362, x_373); lean_dec_ref(x_373); -lean_inc(x_337); -lean_inc(x_346); +lean_inc(x_338); +lean_inc(x_336); x_439 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_439, 0, x_346); -lean_ctor_set(x_439, 1, x_337); +lean_ctor_set(x_439, 0, x_336); +lean_ctor_set(x_439, 1, x_338); lean_ctor_set(x_439, 2, x_438); -lean_inc_n(x_367, 6); -lean_inc(x_346); -x_440 = l_Lean_Syntax_node7(x_346, x_360, x_367, x_367, x_439, x_367, x_367, x_367, x_367); +lean_inc_n(x_370, 6); +lean_inc(x_336); +x_440 = l_Lean_Syntax_node7(x_336, x_365, x_370, x_370, x_439, x_370, x_370, x_370, x_370); x_441 = 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__2_spec__2___closed__5; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_442 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_441); +x_442 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_441); x_443 = 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__2_spec__2___closed__6; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_444 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_443); -lean_inc(x_367); -lean_inc(x_346); -x_445 = l_Lean_Syntax_node1(x_346, x_444, x_367); -lean_inc(x_346); +x_444 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_443); +lean_inc(x_370); +lean_inc(x_336); +x_445 = l_Lean_Syntax_node1(x_336, x_444, x_370); +lean_inc(x_336); lean_ctor_set_tag(x_434, 2); lean_ctor_set(x_434, 1, x_441); -lean_ctor_set(x_434, 0, x_346); -lean_inc(x_367); -lean_inc(x_346); -x_446 = l_Lean_Syntax_node2(x_346, x_343, x_357, x_367); -lean_inc(x_337); -lean_inc(x_346); -x_447 = l_Lean_Syntax_node1(x_346, x_337, x_446); +lean_ctor_set(x_434, 0, x_336); +lean_inc(x_370); +lean_inc(x_336); +x_446 = l_Lean_Syntax_node2(x_336, x_366, x_337, x_370); +lean_inc(x_338); +lean_inc(x_336); +x_447 = l_Lean_Syntax_node1(x_336, x_338, x_446); x_448 = l_Lake_configField___closed__19; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_449 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_448); -lean_inc_ref(x_353); -lean_inc(x_346); +x_449 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_448); +lean_inc_ref(x_361); +lean_inc(x_336); x_450 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_450, 0, x_346); -lean_ctor_set(x_450, 1, x_353); +lean_ctor_set(x_450, 0, x_336); +lean_ctor_set(x_450, 1, x_361); x_451 = 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__4_spec__4___closed__17; x_452 = 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__4_spec__4___closed__18; x_453 = 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__4_spec__4___closed__19; -lean_inc(x_341); -lean_inc(x_356); -x_454 = l_Lean_addMacroScope(x_356, x_453, x_341); -lean_inc_ref(x_355); -x_455 = l_Lean_Name_mkStr2(x_355, x_451); -lean_inc(x_371); +lean_inc(x_369); +lean_inc(x_346); +x_454 = l_Lean_addMacroScope(x_346, x_453, x_369); +lean_inc_ref(x_341); +x_455 = l_Lean_Name_mkStr2(x_341, x_451); +lean_inc(x_339); lean_inc(x_455); x_456 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_456, 0, x_455); -lean_ctor_set(x_456, 1, x_371); +lean_ctor_set(x_456, 1, x_339); x_457 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_457, 0, x_455); -lean_inc(x_352); +lean_inc(x_356); x_458 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_458, 0, x_457); -lean_ctor_set(x_458, 1, x_352); +lean_ctor_set(x_458, 1, x_356); x_459 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_459, 0, x_456); lean_ctor_set(x_459, 1, x_458); -lean_inc(x_346); +lean_inc(x_336); x_460 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_460, 0, x_346); +lean_ctor_set(x_460, 0, x_336); lean_ctor_set(x_460, 1, x_452); lean_ctor_set(x_460, 2, x_454); lean_ctor_set(x_460, 3, x_459); lean_inc(x_1); -lean_inc(x_337); -lean_inc(x_346); -x_461 = l_Lean_Syntax_node2(x_346, x_337, x_1, x_24); -lean_inc(x_346); -x_462 = l_Lean_Syntax_node2(x_346, x_363, x_460, x_461); -lean_inc(x_346); -x_463 = l_Lean_Syntax_node2(x_346, x_349, x_450, x_462); -lean_inc(x_367); -lean_inc(x_346); -x_464 = l_Lean_Syntax_node2(x_346, x_449, x_367, x_463); +lean_inc(x_338); +lean_inc(x_336); +x_461 = l_Lean_Syntax_node2(x_336, x_338, x_1, x_24); +lean_inc(x_336); +x_462 = l_Lean_Syntax_node2(x_336, x_352, x_460, x_461); +lean_inc(x_336); +x_463 = l_Lean_Syntax_node2(x_336, x_353, x_450, x_462); +lean_inc(x_370); +lean_inc(x_336); +x_464 = l_Lean_Syntax_node2(x_336, x_449, x_370, x_463); x_465 = 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__2_spec__2___closed__18; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_466 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_465); -lean_inc_ref(x_372); -lean_inc(x_346); +x_466 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_465); +lean_inc_ref(x_345); +lean_inc(x_336); x_467 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_467, 0, x_346); -lean_ctor_set(x_467, 1, x_372); +lean_ctor_set(x_467, 0, x_336); +lean_ctor_set(x_467, 1, x_345); x_468 = 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__2_spec__2___closed__20; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_469 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_468); +x_469 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_468); x_470 = 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__2_spec__2___closed__21; -lean_inc(x_346); +lean_inc(x_336); x_471 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_471, 0, x_346); +lean_ctor_set(x_471, 0, x_336); lean_ctor_set(x_471, 1, x_470); -lean_inc(x_337); -lean_inc(x_346); -x_472 = l_Lean_Syntax_node1(x_346, x_337, x_345); +lean_inc(x_338); +lean_inc(x_336); +x_472 = l_Lean_Syntax_node1(x_336, x_338, x_371); x_473 = 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__2_spec__2___closed__22; -lean_inc(x_346); +lean_inc(x_336); x_474 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_474, 0, x_346); +lean_ctor_set(x_474, 0, x_336); lean_ctor_set(x_474, 1, x_473); -lean_inc(x_346); -x_475 = l_Lean_Syntax_node3(x_346, x_469, x_471, x_472, x_474); +lean_inc(x_336); +x_475 = l_Lean_Syntax_node3(x_336, x_469, x_471, x_472, x_474); x_476 = 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__2_spec__2___closed__23; x_477 = 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__2_spec__2___closed__24; -x_478 = l_Lean_Name_mkStr4(x_350, x_364, x_476, x_477); -lean_inc_n(x_367, 2); -lean_inc(x_346); -x_479 = l_Lean_Syntax_node2(x_346, x_478, x_367, x_367); -lean_inc(x_367); -lean_inc(x_346); -x_480 = l_Lean_Syntax_node4(x_346, x_466, x_467, x_475, x_479, x_367); -lean_inc(x_346); -x_481 = l_Lean_Syntax_node6(x_346, x_442, x_445, x_434, x_367, x_447, x_464, x_480); -x_482 = l_Lean_Syntax_node2(x_346, x_348, x_440, x_481); -x_483 = lean_array_push(x_336, x_482); +x_478 = l_Lean_Name_mkStr4(x_364, x_343, x_476, x_477); +lean_inc_n(x_370, 2); +lean_inc(x_336); +x_479 = l_Lean_Syntax_node2(x_336, x_478, x_370, x_370); +lean_inc(x_370); +lean_inc(x_336); +x_480 = l_Lean_Syntax_node4(x_336, x_466, x_467, x_475, x_479, x_370); +lean_inc(x_336); +x_481 = l_Lean_Syntax_node6(x_336, x_442, x_445, x_434, x_370, x_447, x_464, x_480); +x_482 = l_Lean_Syntax_node2(x_336, x_358, x_440, x_481); +x_483 = lean_array_push(x_350, x_482); lean_inc(x_436); x_484 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_484, 0, x_436); lean_ctor_set(x_484, 1, x_382); x_485 = 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__2_spec__2___closed__28; x_486 = 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__2_spec__2___closed__29; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_487 = l_Lean_addMacroScope(x_346, x_486, x_369); lean_inc(x_356); -x_487 = l_Lean_addMacroScope(x_356, x_486, x_341); -lean_inc(x_352); lean_inc(x_436); x_488 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_488, 0, x_436); lean_ctor_set(x_488, 1, x_485); lean_ctor_set(x_488, 2, x_487); -lean_ctor_set(x_488, 3, x_352); +lean_ctor_set(x_488, 3, x_356); lean_inc(x_436); x_489 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_489, 0, x_436); -lean_ctor_set(x_489, 1, x_338); -lean_inc(x_337); +lean_ctor_set(x_489, 1, x_360); +lean_inc(x_338); lean_inc(x_436); x_490 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_490, 0, x_436); -lean_ctor_set(x_490, 1, x_337); -lean_ctor_set(x_490, 2, x_370); +lean_ctor_set(x_490, 1, x_338); +lean_ctor_set(x_490, 2, x_362); x_491 = 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__2_spec__2___closed__34; x_492 = 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__2_spec__2___closed__35; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_493 = l_Lean_addMacroScope(x_346, x_492, x_369); lean_inc(x_356); -x_493 = l_Lean_addMacroScope(x_356, x_492, x_341); -lean_inc(x_352); lean_inc(x_436); x_494 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_494, 0, x_436); lean_ctor_set(x_494, 1, x_491); lean_ctor_set(x_494, 2, x_493); -lean_ctor_set(x_494, 3, x_352); +lean_ctor_set(x_494, 3, x_356); lean_inc_ref(x_490); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_436); -x_495 = l_Lean_Syntax_node2(x_436, x_351, x_494, x_490); +x_495 = l_Lean_Syntax_node2(x_436, x_347, x_494, x_490); lean_inc(x_436); x_496 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_496, 0, x_436); -lean_ctor_set(x_496, 1, x_372); +lean_ctor_set(x_496, 1, x_345); lean_inc_ref(x_490); lean_inc_ref(x_496); -lean_inc(x_342); +lean_inc(x_357); lean_inc(x_436); -x_497 = l_Lean_Syntax_node3(x_436, x_342, x_496, x_490, x_361); +x_497 = l_Lean_Syntax_node3(x_436, x_357, x_496, x_490, x_340); lean_inc_ref_n(x_490, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_436); -x_498 = l_Lean_Syntax_node3(x_436, x_337, x_490, x_490, x_497); +x_498 = l_Lean_Syntax_node3(x_436, x_338, x_490, x_490, x_497); lean_inc(x_498); -lean_inc(x_344); +lean_inc(x_349); lean_inc(x_436); -x_499 = l_Lean_Syntax_node2(x_436, x_344, x_495, x_498); +x_499 = l_Lean_Syntax_node2(x_436, x_349, x_495, x_498); x_500 = 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__2_spec__2___closed__38; x_501 = 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__2_spec__2___closed__39; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_502 = l_Lean_addMacroScope(x_346, x_501, x_369); lean_inc(x_356); -x_502 = l_Lean_addMacroScope(x_356, x_501, x_341); -lean_inc(x_352); lean_inc(x_436); x_503 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_503, 0, x_436); lean_ctor_set(x_503, 1, x_500); lean_ctor_set(x_503, 2, x_502); -lean_ctor_set(x_503, 3, x_352); +lean_ctor_set(x_503, 3, x_356); lean_inc_ref(x_490); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_436); -x_504 = l_Lean_Syntax_node2(x_436, x_351, x_503, x_490); -lean_inc(x_344); +x_504 = l_Lean_Syntax_node2(x_436, x_347, x_503, x_490); +lean_inc(x_349); lean_inc(x_436); -x_505 = l_Lean_Syntax_node2(x_436, x_344, x_504, x_498); +x_505 = l_Lean_Syntax_node2(x_436, x_349, x_504, x_498); x_506 = 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__4_spec__4___closed__21; x_507 = 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__4_spec__4___closed__22; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_508 = l_Lean_addMacroScope(x_346, x_507, x_369); lean_inc(x_356); -x_508 = l_Lean_addMacroScope(x_356, x_507, x_341); -lean_inc(x_352); lean_inc(x_436); x_509 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_509, 0, x_436); lean_ctor_set(x_509, 1, x_506); lean_ctor_set(x_509, 2, x_508); -lean_ctor_set(x_509, 3, x_352); +lean_ctor_set(x_509, 3, x_356); lean_inc_ref(x_490); lean_inc(x_436); -x_510 = l_Lean_Syntax_node2(x_436, x_351, x_509, x_490); +x_510 = l_Lean_Syntax_node2(x_436, x_347, x_509, x_490); x_511 = 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__4_spec__4___closed__23; x_512 = 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__4_spec__4___closed__24; -lean_inc(x_341); -lean_inc(x_356); -x_513 = l_Lean_addMacroScope(x_356, x_512, x_341); +lean_inc(x_369); +lean_inc(x_346); +x_513 = l_Lean_addMacroScope(x_346, x_512, x_369); x_514 = 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__2_spec__2___closed__48; -lean_inc(x_371); +lean_inc(x_339); x_515 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_515, 0, x_514); -lean_ctor_set(x_515, 1, x_371); -lean_inc(x_352); +lean_ctor_set(x_515, 1, x_339); +lean_inc(x_356); x_516 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_516, 0, x_515); -lean_ctor_set(x_516, 1, x_352); +lean_ctor_set(x_516, 1, x_356); lean_inc(x_436); x_517 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_517, 0, x_436); @@ -6237,40 +6237,40 @@ lean_ctor_set(x_517, 2, x_513); lean_ctor_set(x_517, 3, x_516); lean_inc_ref(x_490); lean_inc(x_436); -x_518 = l_Lean_Syntax_node3(x_436, x_342, x_496, x_490, x_517); +x_518 = l_Lean_Syntax_node3(x_436, x_357, x_496, x_490, x_517); lean_inc_ref_n(x_490, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_436); -x_519 = l_Lean_Syntax_node3(x_436, x_337, x_490, x_490, x_518); +x_519 = l_Lean_Syntax_node3(x_436, x_338, x_490, x_490, x_518); lean_inc(x_436); -x_520 = l_Lean_Syntax_node2(x_436, x_344, x_510, x_519); +x_520 = l_Lean_Syntax_node2(x_436, x_349, x_510, x_519); lean_inc_ref_n(x_490, 3); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_436); -x_521 = l_Lean_Syntax_node6(x_436, x_337, x_499, x_490, x_505, x_490, x_520, x_490); +x_521 = l_Lean_Syntax_node6(x_436, x_338, x_499, x_490, x_505, x_490, x_520, x_490); lean_inc(x_436); -x_522 = l_Lean_Syntax_node1(x_436, x_339, x_521); +x_522 = l_Lean_Syntax_node1(x_436, x_363, x_521); lean_inc_ref(x_490); lean_inc(x_436); -x_523 = l_Lean_Syntax_node1(x_436, x_358, x_490); +x_523 = l_Lean_Syntax_node1(x_436, x_355, x_490); lean_inc(x_436); x_524 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_524, 0, x_436); -lean_ctor_set(x_524, 1, x_353); +lean_ctor_set(x_524, 1, x_361); x_525 = 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__2_spec__2___closed__1; x_526 = 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__2_spec__2___closed__2; x_527 = 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__2_spec__2___closed__3; -x_528 = l_Lean_addMacroScope(x_356, x_527, x_341); -x_529 = l_Lean_Name_mkStr2(x_355, x_525); +x_528 = l_Lean_addMacroScope(x_346, x_527, x_369); +x_529 = l_Lean_Name_mkStr2(x_341, x_525); lean_inc(x_529); x_530 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_530, 0, x_529); -lean_ctor_set(x_530, 1, x_371); +lean_ctor_set(x_530, 1, x_339); x_531 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_531, 0, x_529); x_532 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_532, 0, x_531); -lean_ctor_set(x_532, 1, x_352); +lean_ctor_set(x_532, 1, x_356); x_533 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_533, 0, x_530); lean_ctor_set(x_533, 1, x_532); @@ -6280,17 +6280,17 @@ lean_ctor_set(x_534, 0, x_436); lean_ctor_set(x_534, 1, x_526); lean_ctor_set(x_534, 2, x_528); lean_ctor_set(x_534, 3, x_533); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_436); -x_535 = l_Lean_Syntax_node2(x_436, x_337, x_524, x_534); +x_535 = l_Lean_Syntax_node2(x_436, x_338, x_524, x_534); lean_inc(x_436); x_536 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_536, 0, x_436); -lean_ctor_set(x_536, 1, x_354); +lean_ctor_set(x_536, 1, x_342); lean_inc(x_436); -x_537 = l_Lean_Syntax_node6(x_436, x_366, x_489, x_490, x_522, x_523, x_535, x_536); +x_537 = l_Lean_Syntax_node6(x_436, x_351, x_489, x_490, x_522, x_523, x_535, x_536); lean_inc(x_436); -x_538 = l_Lean_Syntax_node1(x_436, x_337, x_537); +x_538 = l_Lean_Syntax_node1(x_436, x_338, x_537); x_539 = l_Lean_Syntax_node4(x_436, x_381, x_431, x_484, x_488, x_538); x_540 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_540, 0, x_483); @@ -6307,238 +6307,238 @@ x_542 = lean_ctor_get(x_434, 1); lean_inc(x_542); lean_inc(x_541); lean_dec(x_434); -lean_inc_ref(x_370); -x_543 = l_Array_append___redArg(x_370, x_373); +lean_inc_ref(x_362); +x_543 = l_Array_append___redArg(x_362, x_373); lean_dec_ref(x_373); -lean_inc(x_337); -lean_inc(x_346); +lean_inc(x_338); +lean_inc(x_336); x_544 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_544, 0, x_346); -lean_ctor_set(x_544, 1, x_337); +lean_ctor_set(x_544, 0, x_336); +lean_ctor_set(x_544, 1, x_338); lean_ctor_set(x_544, 2, x_543); -lean_inc_n(x_367, 6); -lean_inc(x_346); -x_545 = l_Lean_Syntax_node7(x_346, x_360, x_367, x_367, x_544, x_367, x_367, x_367, x_367); +lean_inc_n(x_370, 6); +lean_inc(x_336); +x_545 = l_Lean_Syntax_node7(x_336, x_365, x_370, x_370, x_544, x_370, x_370, x_370, x_370); x_546 = 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__2_spec__2___closed__5; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_547 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_546); +x_547 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_546); x_548 = 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__2_spec__2___closed__6; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_549 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_548); -lean_inc(x_367); -lean_inc(x_346); -x_550 = l_Lean_Syntax_node1(x_346, x_549, x_367); -lean_inc(x_346); +x_549 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_548); +lean_inc(x_370); +lean_inc(x_336); +x_550 = l_Lean_Syntax_node1(x_336, x_549, x_370); +lean_inc(x_336); x_551 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_551, 0, x_346); +lean_ctor_set(x_551, 0, x_336); lean_ctor_set(x_551, 1, x_546); -lean_inc(x_367); -lean_inc(x_346); -x_552 = l_Lean_Syntax_node2(x_346, x_343, x_357, x_367); -lean_inc(x_337); -lean_inc(x_346); -x_553 = l_Lean_Syntax_node1(x_346, x_337, x_552); +lean_inc(x_370); +lean_inc(x_336); +x_552 = l_Lean_Syntax_node2(x_336, x_366, x_337, x_370); +lean_inc(x_338); +lean_inc(x_336); +x_553 = l_Lean_Syntax_node1(x_336, x_338, x_552); x_554 = l_Lake_configField___closed__19; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_555 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_554); -lean_inc_ref(x_353); -lean_inc(x_346); +x_555 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_554); +lean_inc_ref(x_361); +lean_inc(x_336); x_556 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_556, 0, x_346); -lean_ctor_set(x_556, 1, x_353); +lean_ctor_set(x_556, 0, x_336); +lean_ctor_set(x_556, 1, x_361); x_557 = 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__4_spec__4___closed__17; x_558 = 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__4_spec__4___closed__18; x_559 = 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__4_spec__4___closed__19; -lean_inc(x_341); -lean_inc(x_356); -x_560 = l_Lean_addMacroScope(x_356, x_559, x_341); -lean_inc_ref(x_355); -x_561 = l_Lean_Name_mkStr2(x_355, x_557); -lean_inc(x_371); +lean_inc(x_369); +lean_inc(x_346); +x_560 = l_Lean_addMacroScope(x_346, x_559, x_369); +lean_inc_ref(x_341); +x_561 = l_Lean_Name_mkStr2(x_341, x_557); +lean_inc(x_339); lean_inc(x_561); x_562 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_562, 0, x_561); -lean_ctor_set(x_562, 1, x_371); +lean_ctor_set(x_562, 1, x_339); x_563 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_563, 0, x_561); -lean_inc(x_352); +lean_inc(x_356); x_564 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_564, 0, x_563); -lean_ctor_set(x_564, 1, x_352); +lean_ctor_set(x_564, 1, x_356); x_565 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_565, 0, x_562); lean_ctor_set(x_565, 1, x_564); -lean_inc(x_346); +lean_inc(x_336); x_566 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_566, 0, x_346); +lean_ctor_set(x_566, 0, x_336); lean_ctor_set(x_566, 1, x_558); lean_ctor_set(x_566, 2, x_560); lean_ctor_set(x_566, 3, x_565); lean_inc(x_1); -lean_inc(x_337); -lean_inc(x_346); -x_567 = l_Lean_Syntax_node2(x_346, x_337, x_1, x_24); -lean_inc(x_346); -x_568 = l_Lean_Syntax_node2(x_346, x_363, x_566, x_567); -lean_inc(x_346); -x_569 = l_Lean_Syntax_node2(x_346, x_349, x_556, x_568); -lean_inc(x_367); -lean_inc(x_346); -x_570 = l_Lean_Syntax_node2(x_346, x_555, x_367, x_569); +lean_inc(x_338); +lean_inc(x_336); +x_567 = l_Lean_Syntax_node2(x_336, x_338, x_1, x_24); +lean_inc(x_336); +x_568 = l_Lean_Syntax_node2(x_336, x_352, x_566, x_567); +lean_inc(x_336); +x_569 = l_Lean_Syntax_node2(x_336, x_353, x_556, x_568); +lean_inc(x_370); +lean_inc(x_336); +x_570 = l_Lean_Syntax_node2(x_336, x_555, x_370, x_569); x_571 = 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__2_spec__2___closed__18; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_572 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_571); -lean_inc_ref(x_372); -lean_inc(x_346); +x_572 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_571); +lean_inc_ref(x_345); +lean_inc(x_336); x_573 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_573, 0, x_346); -lean_ctor_set(x_573, 1, x_372); +lean_ctor_set(x_573, 0, x_336); +lean_ctor_set(x_573, 1, x_345); x_574 = 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__2_spec__2___closed__20; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_575 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_574); +x_575 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_574); x_576 = 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__2_spec__2___closed__21; -lean_inc(x_346); +lean_inc(x_336); x_577 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_577, 0, x_346); +lean_ctor_set(x_577, 0, x_336); lean_ctor_set(x_577, 1, x_576); -lean_inc(x_337); -lean_inc(x_346); -x_578 = l_Lean_Syntax_node1(x_346, x_337, x_345); +lean_inc(x_338); +lean_inc(x_336); +x_578 = l_Lean_Syntax_node1(x_336, x_338, x_371); x_579 = 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__2_spec__2___closed__22; -lean_inc(x_346); +lean_inc(x_336); x_580 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_580, 0, x_346); +lean_ctor_set(x_580, 0, x_336); lean_ctor_set(x_580, 1, x_579); -lean_inc(x_346); -x_581 = l_Lean_Syntax_node3(x_346, x_575, x_577, x_578, x_580); +lean_inc(x_336); +x_581 = l_Lean_Syntax_node3(x_336, x_575, x_577, x_578, x_580); x_582 = 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__2_spec__2___closed__23; x_583 = 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__2_spec__2___closed__24; -x_584 = l_Lean_Name_mkStr4(x_350, x_364, x_582, x_583); -lean_inc_n(x_367, 2); -lean_inc(x_346); -x_585 = l_Lean_Syntax_node2(x_346, x_584, x_367, x_367); -lean_inc(x_367); -lean_inc(x_346); -x_586 = l_Lean_Syntax_node4(x_346, x_572, x_573, x_581, x_585, x_367); -lean_inc(x_346); -x_587 = l_Lean_Syntax_node6(x_346, x_547, x_550, x_551, x_367, x_553, x_570, x_586); -x_588 = l_Lean_Syntax_node2(x_346, x_348, x_545, x_587); -x_589 = lean_array_push(x_336, x_588); +x_584 = l_Lean_Name_mkStr4(x_364, x_343, x_582, x_583); +lean_inc_n(x_370, 2); +lean_inc(x_336); +x_585 = l_Lean_Syntax_node2(x_336, x_584, x_370, x_370); +lean_inc(x_370); +lean_inc(x_336); +x_586 = l_Lean_Syntax_node4(x_336, x_572, x_573, x_581, x_585, x_370); +lean_inc(x_336); +x_587 = l_Lean_Syntax_node6(x_336, x_547, x_550, x_551, x_370, x_553, x_570, x_586); +x_588 = l_Lean_Syntax_node2(x_336, x_358, x_545, x_587); +x_589 = lean_array_push(x_350, x_588); lean_inc(x_541); x_590 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_590, 0, x_541); lean_ctor_set(x_590, 1, x_382); x_591 = 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__2_spec__2___closed__28; x_592 = 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__2_spec__2___closed__29; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_593 = l_Lean_addMacroScope(x_346, x_592, x_369); lean_inc(x_356); -x_593 = l_Lean_addMacroScope(x_356, x_592, x_341); -lean_inc(x_352); lean_inc(x_541); x_594 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_594, 0, x_541); lean_ctor_set(x_594, 1, x_591); lean_ctor_set(x_594, 2, x_593); -lean_ctor_set(x_594, 3, x_352); +lean_ctor_set(x_594, 3, x_356); lean_inc(x_541); x_595 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_595, 0, x_541); -lean_ctor_set(x_595, 1, x_338); -lean_inc(x_337); +lean_ctor_set(x_595, 1, x_360); +lean_inc(x_338); lean_inc(x_541); x_596 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_596, 0, x_541); -lean_ctor_set(x_596, 1, x_337); -lean_ctor_set(x_596, 2, x_370); +lean_ctor_set(x_596, 1, x_338); +lean_ctor_set(x_596, 2, x_362); x_597 = 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__2_spec__2___closed__34; x_598 = 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__2_spec__2___closed__35; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_599 = l_Lean_addMacroScope(x_346, x_598, x_369); lean_inc(x_356); -x_599 = l_Lean_addMacroScope(x_356, x_598, x_341); -lean_inc(x_352); lean_inc(x_541); x_600 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_600, 0, x_541); lean_ctor_set(x_600, 1, x_597); lean_ctor_set(x_600, 2, x_599); -lean_ctor_set(x_600, 3, x_352); +lean_ctor_set(x_600, 3, x_356); lean_inc_ref(x_596); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_541); -x_601 = l_Lean_Syntax_node2(x_541, x_351, x_600, x_596); +x_601 = l_Lean_Syntax_node2(x_541, x_347, x_600, x_596); lean_inc(x_541); x_602 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_602, 0, x_541); -lean_ctor_set(x_602, 1, x_372); +lean_ctor_set(x_602, 1, x_345); lean_inc_ref(x_596); lean_inc_ref(x_602); -lean_inc(x_342); +lean_inc(x_357); lean_inc(x_541); -x_603 = l_Lean_Syntax_node3(x_541, x_342, x_602, x_596, x_361); +x_603 = l_Lean_Syntax_node3(x_541, x_357, x_602, x_596, x_340); lean_inc_ref_n(x_596, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_541); -x_604 = l_Lean_Syntax_node3(x_541, x_337, x_596, x_596, x_603); +x_604 = l_Lean_Syntax_node3(x_541, x_338, x_596, x_596, x_603); lean_inc(x_604); -lean_inc(x_344); +lean_inc(x_349); lean_inc(x_541); -x_605 = l_Lean_Syntax_node2(x_541, x_344, x_601, x_604); +x_605 = l_Lean_Syntax_node2(x_541, x_349, x_601, x_604); x_606 = 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__2_spec__2___closed__38; x_607 = 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__2_spec__2___closed__39; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_608 = l_Lean_addMacroScope(x_346, x_607, x_369); lean_inc(x_356); -x_608 = l_Lean_addMacroScope(x_356, x_607, x_341); -lean_inc(x_352); lean_inc(x_541); x_609 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_609, 0, x_541); lean_ctor_set(x_609, 1, x_606); lean_ctor_set(x_609, 2, x_608); -lean_ctor_set(x_609, 3, x_352); +lean_ctor_set(x_609, 3, x_356); lean_inc_ref(x_596); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_541); -x_610 = l_Lean_Syntax_node2(x_541, x_351, x_609, x_596); -lean_inc(x_344); +x_610 = l_Lean_Syntax_node2(x_541, x_347, x_609, x_596); +lean_inc(x_349); lean_inc(x_541); -x_611 = l_Lean_Syntax_node2(x_541, x_344, x_610, x_604); +x_611 = l_Lean_Syntax_node2(x_541, x_349, x_610, x_604); x_612 = 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__4_spec__4___closed__21; x_613 = 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__4_spec__4___closed__22; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_614 = l_Lean_addMacroScope(x_346, x_613, x_369); lean_inc(x_356); -x_614 = l_Lean_addMacroScope(x_356, x_613, x_341); -lean_inc(x_352); lean_inc(x_541); x_615 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_615, 0, x_541); lean_ctor_set(x_615, 1, x_612); lean_ctor_set(x_615, 2, x_614); -lean_ctor_set(x_615, 3, x_352); +lean_ctor_set(x_615, 3, x_356); lean_inc_ref(x_596); lean_inc(x_541); -x_616 = l_Lean_Syntax_node2(x_541, x_351, x_615, x_596); +x_616 = l_Lean_Syntax_node2(x_541, x_347, x_615, x_596); x_617 = 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__4_spec__4___closed__23; x_618 = 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__4_spec__4___closed__24; -lean_inc(x_341); -lean_inc(x_356); -x_619 = l_Lean_addMacroScope(x_356, x_618, x_341); +lean_inc(x_369); +lean_inc(x_346); +x_619 = l_Lean_addMacroScope(x_346, x_618, x_369); x_620 = 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__2_spec__2___closed__48; -lean_inc(x_371); +lean_inc(x_339); x_621 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_621, 0, x_620); -lean_ctor_set(x_621, 1, x_371); -lean_inc(x_352); +lean_ctor_set(x_621, 1, x_339); +lean_inc(x_356); x_622 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_622, 0, x_621); -lean_ctor_set(x_622, 1, x_352); +lean_ctor_set(x_622, 1, x_356); lean_inc(x_541); x_623 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_623, 0, x_541); @@ -6547,40 +6547,40 @@ lean_ctor_set(x_623, 2, x_619); lean_ctor_set(x_623, 3, x_622); lean_inc_ref(x_596); lean_inc(x_541); -x_624 = l_Lean_Syntax_node3(x_541, x_342, x_602, x_596, x_623); +x_624 = l_Lean_Syntax_node3(x_541, x_357, x_602, x_596, x_623); lean_inc_ref_n(x_596, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_541); -x_625 = l_Lean_Syntax_node3(x_541, x_337, x_596, x_596, x_624); +x_625 = l_Lean_Syntax_node3(x_541, x_338, x_596, x_596, x_624); lean_inc(x_541); -x_626 = l_Lean_Syntax_node2(x_541, x_344, x_616, x_625); +x_626 = l_Lean_Syntax_node2(x_541, x_349, x_616, x_625); lean_inc_ref_n(x_596, 3); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_541); -x_627 = l_Lean_Syntax_node6(x_541, x_337, x_605, x_596, x_611, x_596, x_626, x_596); +x_627 = l_Lean_Syntax_node6(x_541, x_338, x_605, x_596, x_611, x_596, x_626, x_596); lean_inc(x_541); -x_628 = l_Lean_Syntax_node1(x_541, x_339, x_627); +x_628 = l_Lean_Syntax_node1(x_541, x_363, x_627); lean_inc_ref(x_596); lean_inc(x_541); -x_629 = l_Lean_Syntax_node1(x_541, x_358, x_596); +x_629 = l_Lean_Syntax_node1(x_541, x_355, x_596); lean_inc(x_541); x_630 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_630, 0, x_541); -lean_ctor_set(x_630, 1, x_353); +lean_ctor_set(x_630, 1, x_361); x_631 = 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__2_spec__2___closed__1; x_632 = 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__2_spec__2___closed__2; x_633 = 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__2_spec__2___closed__3; -x_634 = l_Lean_addMacroScope(x_356, x_633, x_341); -x_635 = l_Lean_Name_mkStr2(x_355, x_631); +x_634 = l_Lean_addMacroScope(x_346, x_633, x_369); +x_635 = l_Lean_Name_mkStr2(x_341, x_631); lean_inc(x_635); x_636 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_636, 0, x_635); -lean_ctor_set(x_636, 1, x_371); +lean_ctor_set(x_636, 1, x_339); x_637 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_637, 0, x_635); x_638 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_638, 0, x_637); -lean_ctor_set(x_638, 1, x_352); +lean_ctor_set(x_638, 1, x_356); x_639 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_639, 0, x_636); lean_ctor_set(x_639, 1, x_638); @@ -6590,17 +6590,17 @@ lean_ctor_set(x_640, 0, x_541); lean_ctor_set(x_640, 1, x_632); lean_ctor_set(x_640, 2, x_634); lean_ctor_set(x_640, 3, x_639); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_541); -x_641 = l_Lean_Syntax_node2(x_541, x_337, x_630, x_640); +x_641 = l_Lean_Syntax_node2(x_541, x_338, x_630, x_640); lean_inc(x_541); x_642 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_642, 0, x_541); -lean_ctor_set(x_642, 1, x_354); +lean_ctor_set(x_642, 1, x_342); lean_inc(x_541); -x_643 = l_Lean_Syntax_node6(x_541, x_366, x_595, x_596, x_628, x_629, x_641, x_642); +x_643 = l_Lean_Syntax_node6(x_541, x_351, x_595, x_596, x_628, x_629, x_641, x_642); lean_inc(x_541); -x_644 = l_Lean_Syntax_node1(x_541, x_337, x_643); +x_644 = l_Lean_Syntax_node1(x_541, x_338, x_643); x_645 = l_Lean_Syntax_node4(x_541, x_381, x_431, x_590, x_594, x_644); x_646 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_646, 0, x_589); @@ -6619,10 +6619,10 @@ lean_inc(x_648); lean_inc(x_647); lean_dec(x_376); x_649 = 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__2_spec__2___closed__25; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_650 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_649); +x_650 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_649); x_651 = 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__2_spec__2___closed__26; lean_inc(x_647); x_652 = lean_alloc_ctor(2, 2, 0); @@ -6630,26 +6630,26 @@ lean_ctor_set(x_652, 0, x_647); lean_ctor_set(x_652, 1, x_651); x_653 = 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__4_spec__4___closed__1; x_654 = 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__4_spec__4___closed__2; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_655 = l_Lean_addMacroScope(x_346, x_654, x_369); lean_inc(x_356); -x_655 = l_Lean_addMacroScope(x_356, x_654, x_341); -lean_inc(x_352); lean_inc(x_647); x_656 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_656, 0, x_647); lean_ctor_set(x_656, 1, x_653); lean_ctor_set(x_656, 2, x_655); -lean_ctor_set(x_656, 3, x_352); +lean_ctor_set(x_656, 3, x_356); x_657 = 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__4_spec__4___closed__3; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_658 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_657); +x_658 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_657); x_659 = 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__4_spec__4___closed__4; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_660 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_659); +x_660 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_659); x_661 = 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__4_spec__4___closed__5; lean_inc(x_647); x_662 = lean_alloc_ctor(2, 2, 0); @@ -6658,37 +6658,37 @@ lean_ctor_set(x_662, 1, x_661); x_663 = 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__4_spec__4___closed__7; x_664 = 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__4_spec__4___closed__9; x_665 = lean_box(0); -lean_inc(x_341); -lean_inc(x_356); -x_666 = l_Lean_addMacroScope(x_356, x_665, x_341); -lean_inc_ref(x_355); -x_667 = l_Lean_Name_mkStr1(x_355); +lean_inc(x_369); +lean_inc(x_346); +x_666 = l_Lean_addMacroScope(x_346, x_665, x_369); +lean_inc_ref(x_341); +x_667 = l_Lean_Name_mkStr1(x_341); x_668 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_668, 0, x_667); -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_669 = l_Lean_Name_mkStr3(x_350, x_364, x_362); +x_669 = l_Lean_Name_mkStr3(x_364, x_343, x_367); x_670 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_670, 0, x_669); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_671 = l_Lean_Name_mkStr2(x_350, x_364); +x_671 = l_Lean_Name_mkStr2(x_364, x_343); x_672 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_672, 0, x_671); x_673 = 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__4_spec__4___closed__10; -lean_inc_ref(x_350); -x_674 = l_Lean_Name_mkStr2(x_350, x_673); +lean_inc_ref(x_364); +x_674 = l_Lean_Name_mkStr2(x_364, x_673); x_675 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_675, 0, x_674); -lean_inc_ref(x_350); -x_676 = l_Lean_Name_mkStr1(x_350); +lean_inc_ref(x_364); +x_676 = l_Lean_Name_mkStr1(x_364); x_677 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_677, 0, x_676); -lean_inc(x_352); +lean_inc(x_356); x_678 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_678, 0, x_677); -lean_ctor_set(x_678, 1, x_352); +lean_ctor_set(x_678, 1, x_356); x_679 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_679, 0, x_675); lean_ctor_set(x_679, 1, x_678); @@ -6715,19 +6715,19 @@ x_686 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__ x_687 = 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__4_spec__4___closed__13; x_688 = 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__4_spec__4___closed__14; x_689 = 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__4_spec__4___closed__15; -lean_inc(x_341); -lean_inc(x_356); -x_690 = l_Lean_addMacroScope(x_356, x_689, x_341); -lean_inc_ref(x_355); -x_691 = l_Lean_Name_mkStr3(x_355, x_687, x_688); -lean_inc(x_371); +lean_inc(x_369); +lean_inc(x_346); +x_690 = l_Lean_addMacroScope(x_346, x_689, x_369); +lean_inc_ref(x_341); +x_691 = l_Lean_Name_mkStr3(x_341, x_687, x_688); +lean_inc(x_339); x_692 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_692, 0, x_691); -lean_ctor_set(x_692, 1, x_371); -lean_inc(x_352); +lean_ctor_set(x_692, 1, x_339); +lean_inc(x_356); x_693 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_693, 0, x_692); -lean_ctor_set(x_693, 1, x_352); +lean_ctor_set(x_693, 1, x_356); lean_inc(x_647); x_694 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_694, 0, x_647); @@ -6735,12 +6735,12 @@ lean_ctor_set(x_694, 1, x_686); lean_ctor_set(x_694, 2, x_690); lean_ctor_set(x_694, 3, x_693); lean_inc(x_24); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_647); -x_695 = l_Lean_Syntax_node1(x_647, x_337, x_24); -lean_inc(x_363); +x_695 = l_Lean_Syntax_node1(x_647, x_338, x_24); +lean_inc(x_352); lean_inc(x_647); -x_696 = l_Lean_Syntax_node2(x_647, x_363, x_694, x_695); +x_696 = l_Lean_Syntax_node2(x_647, x_352, x_694, x_695); x_697 = 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__4_spec__4___closed__16; lean_inc(x_647); x_698 = lean_alloc_ctor(2, 2, 0); @@ -6748,22 +6748,22 @@ lean_ctor_set(x_698, 0, x_647); lean_ctor_set(x_698, 1, x_697); lean_inc(x_647); x_699 = l_Lean_Syntax_node3(x_647, x_658, x_685, x_696, x_698); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_647); -x_700 = l_Lean_Syntax_node1(x_647, x_337, x_699); +x_700 = l_Lean_Syntax_node1(x_647, x_338, x_699); lean_inc(x_650); x_701 = l_Lean_Syntax_node4(x_647, x_650, x_19, x_652, x_656, x_700); -x_702 = l_Lean_replaceRef(x_701, x_347); -lean_dec(x_347); +x_702 = l_Lean_replaceRef(x_701, x_344); +lean_dec(x_344); lean_inc(x_702); -lean_inc(x_341); -lean_inc(x_356); +lean_inc(x_369); +lean_inc(x_346); x_703 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_703, 0, x_365); -lean_ctor_set(x_703, 1, x_356); -lean_ctor_set(x_703, 2, x_341); -lean_ctor_set(x_703, 3, x_369); -lean_ctor_set(x_703, 4, x_368); +lean_ctor_set(x_703, 0, x_372); +lean_ctor_set(x_703, 1, x_346); +lean_ctor_set(x_703, 2, x_369); +lean_ctor_set(x_703, 3, x_354); +lean_ctor_set(x_703, 4, x_348); lean_ctor_set(x_703, 5, x_702); x_704 = 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__2_spec__2___lam__1(x_17, x_702, x_703, x_648); lean_dec_ref(x_703); @@ -6780,243 +6780,243 @@ if (lean_is_exclusive(x_704)) { lean_dec_ref(x_704); x_707 = lean_box(0); } -lean_inc_ref(x_370); -x_708 = l_Array_append___redArg(x_370, x_373); +lean_inc_ref(x_362); +x_708 = l_Array_append___redArg(x_362, x_373); lean_dec_ref(x_373); -lean_inc(x_337); -lean_inc(x_346); +lean_inc(x_338); +lean_inc(x_336); x_709 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_709, 0, x_346); -lean_ctor_set(x_709, 1, x_337); +lean_ctor_set(x_709, 0, x_336); +lean_ctor_set(x_709, 1, x_338); lean_ctor_set(x_709, 2, x_708); -lean_inc_n(x_367, 6); -lean_inc(x_346); -x_710 = l_Lean_Syntax_node7(x_346, x_360, x_367, x_367, x_709, x_367, x_367, x_367, x_367); +lean_inc_n(x_370, 6); +lean_inc(x_336); +x_710 = l_Lean_Syntax_node7(x_336, x_365, x_370, x_370, x_709, x_370, x_370, x_370, x_370); x_711 = 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__2_spec__2___closed__5; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_712 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_711); +x_712 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_711); x_713 = 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__2_spec__2___closed__6; -lean_inc_ref(x_340); +lean_inc_ref(x_359); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_714 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_713); -lean_inc(x_367); -lean_inc(x_346); -x_715 = l_Lean_Syntax_node1(x_346, x_714, x_367); -lean_inc(x_346); +x_714 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_713); +lean_inc(x_370); +lean_inc(x_336); +x_715 = l_Lean_Syntax_node1(x_336, x_714, x_370); +lean_inc(x_336); if (lean_is_scalar(x_707)) { x_716 = lean_alloc_ctor(2, 2, 0); } else { x_716 = x_707; lean_ctor_set_tag(x_716, 2); } -lean_ctor_set(x_716, 0, x_346); +lean_ctor_set(x_716, 0, x_336); lean_ctor_set(x_716, 1, x_711); -lean_inc(x_367); -lean_inc(x_346); -x_717 = l_Lean_Syntax_node2(x_346, x_343, x_357, x_367); -lean_inc(x_337); -lean_inc(x_346); -x_718 = l_Lean_Syntax_node1(x_346, x_337, x_717); +lean_inc(x_370); +lean_inc(x_336); +x_717 = l_Lean_Syntax_node2(x_336, x_366, x_337, x_370); +lean_inc(x_338); +lean_inc(x_336); +x_718 = l_Lean_Syntax_node1(x_336, x_338, x_717); x_719 = l_Lake_configField___closed__19; -lean_inc_ref(x_362); +lean_inc_ref(x_367); +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_720 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_719); -lean_inc_ref(x_353); -lean_inc(x_346); +x_720 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_719); +lean_inc_ref(x_361); +lean_inc(x_336); x_721 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_721, 0, x_346); -lean_ctor_set(x_721, 1, x_353); +lean_ctor_set(x_721, 0, x_336); +lean_ctor_set(x_721, 1, x_361); x_722 = 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__4_spec__4___closed__17; x_723 = 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__4_spec__4___closed__18; x_724 = 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__4_spec__4___closed__19; -lean_inc(x_341); -lean_inc(x_356); -x_725 = l_Lean_addMacroScope(x_356, x_724, x_341); -lean_inc_ref(x_355); -x_726 = l_Lean_Name_mkStr2(x_355, x_722); -lean_inc(x_371); +lean_inc(x_369); +lean_inc(x_346); +x_725 = l_Lean_addMacroScope(x_346, x_724, x_369); +lean_inc_ref(x_341); +x_726 = l_Lean_Name_mkStr2(x_341, x_722); +lean_inc(x_339); lean_inc(x_726); x_727 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_727, 0, x_726); -lean_ctor_set(x_727, 1, x_371); +lean_ctor_set(x_727, 1, x_339); x_728 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_728, 0, x_726); -lean_inc(x_352); +lean_inc(x_356); x_729 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_729, 0, x_728); -lean_ctor_set(x_729, 1, x_352); +lean_ctor_set(x_729, 1, x_356); x_730 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_730, 0, x_727); lean_ctor_set(x_730, 1, x_729); -lean_inc(x_346); +lean_inc(x_336); x_731 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_731, 0, x_346); +lean_ctor_set(x_731, 0, x_336); lean_ctor_set(x_731, 1, x_723); lean_ctor_set(x_731, 2, x_725); lean_ctor_set(x_731, 3, x_730); lean_inc(x_1); -lean_inc(x_337); -lean_inc(x_346); -x_732 = l_Lean_Syntax_node2(x_346, x_337, x_1, x_24); -lean_inc(x_346); -x_733 = l_Lean_Syntax_node2(x_346, x_363, x_731, x_732); -lean_inc(x_346); -x_734 = l_Lean_Syntax_node2(x_346, x_349, x_721, x_733); -lean_inc(x_367); -lean_inc(x_346); -x_735 = l_Lean_Syntax_node2(x_346, x_720, x_367, x_734); +lean_inc(x_338); +lean_inc(x_336); +x_732 = l_Lean_Syntax_node2(x_336, x_338, x_1, x_24); +lean_inc(x_336); +x_733 = l_Lean_Syntax_node2(x_336, x_352, x_731, x_732); +lean_inc(x_336); +x_734 = l_Lean_Syntax_node2(x_336, x_353, x_721, x_733); +lean_inc(x_370); +lean_inc(x_336); +x_735 = l_Lean_Syntax_node2(x_336, x_720, x_370, x_734); x_736 = 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__2_spec__2___closed__18; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_737 = l_Lean_Name_mkStr4(x_350, x_364, x_362, x_736); -lean_inc_ref(x_372); -lean_inc(x_346); +x_737 = l_Lean_Name_mkStr4(x_364, x_343, x_367, x_736); +lean_inc_ref(x_345); +lean_inc(x_336); x_738 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_738, 0, x_346); -lean_ctor_set(x_738, 1, x_372); +lean_ctor_set(x_738, 0, x_336); +lean_ctor_set(x_738, 1, x_345); x_739 = 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__2_spec__2___closed__20; +lean_inc_ref(x_343); lean_inc_ref(x_364); -lean_inc_ref(x_350); -x_740 = l_Lean_Name_mkStr4(x_350, x_364, x_340, x_739); +x_740 = l_Lean_Name_mkStr4(x_364, x_343, x_359, x_739); x_741 = 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__2_spec__2___closed__21; -lean_inc(x_346); +lean_inc(x_336); x_742 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_742, 0, x_346); +lean_ctor_set(x_742, 0, x_336); lean_ctor_set(x_742, 1, x_741); -lean_inc(x_337); -lean_inc(x_346); -x_743 = l_Lean_Syntax_node1(x_346, x_337, x_345); +lean_inc(x_338); +lean_inc(x_336); +x_743 = l_Lean_Syntax_node1(x_336, x_338, x_371); x_744 = 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__2_spec__2___closed__22; -lean_inc(x_346); +lean_inc(x_336); x_745 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_745, 0, x_346); +lean_ctor_set(x_745, 0, x_336); lean_ctor_set(x_745, 1, x_744); -lean_inc(x_346); -x_746 = l_Lean_Syntax_node3(x_346, x_740, x_742, x_743, x_745); +lean_inc(x_336); +x_746 = l_Lean_Syntax_node3(x_336, x_740, x_742, x_743, x_745); x_747 = 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__2_spec__2___closed__23; x_748 = 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__2_spec__2___closed__24; -x_749 = l_Lean_Name_mkStr4(x_350, x_364, x_747, x_748); -lean_inc_n(x_367, 2); -lean_inc(x_346); -x_750 = l_Lean_Syntax_node2(x_346, x_749, x_367, x_367); -lean_inc(x_367); -lean_inc(x_346); -x_751 = l_Lean_Syntax_node4(x_346, x_737, x_738, x_746, x_750, x_367); -lean_inc(x_346); -x_752 = l_Lean_Syntax_node6(x_346, x_712, x_715, x_716, x_367, x_718, x_735, x_751); -x_753 = l_Lean_Syntax_node2(x_346, x_348, x_710, x_752); -x_754 = lean_array_push(x_336, x_753); +x_749 = l_Lean_Name_mkStr4(x_364, x_343, x_747, x_748); +lean_inc_n(x_370, 2); +lean_inc(x_336); +x_750 = l_Lean_Syntax_node2(x_336, x_749, x_370, x_370); +lean_inc(x_370); +lean_inc(x_336); +x_751 = l_Lean_Syntax_node4(x_336, x_737, x_738, x_746, x_750, x_370); +lean_inc(x_336); +x_752 = l_Lean_Syntax_node6(x_336, x_712, x_715, x_716, x_370, x_718, x_735, x_751); +x_753 = l_Lean_Syntax_node2(x_336, x_358, x_710, x_752); +x_754 = lean_array_push(x_350, x_753); lean_inc(x_705); x_755 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_755, 0, x_705); lean_ctor_set(x_755, 1, x_651); x_756 = 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__2_spec__2___closed__28; x_757 = 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__2_spec__2___closed__29; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_758 = l_Lean_addMacroScope(x_346, x_757, x_369); lean_inc(x_356); -x_758 = l_Lean_addMacroScope(x_356, x_757, x_341); -lean_inc(x_352); lean_inc(x_705); x_759 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_759, 0, x_705); lean_ctor_set(x_759, 1, x_756); lean_ctor_set(x_759, 2, x_758); -lean_ctor_set(x_759, 3, x_352); +lean_ctor_set(x_759, 3, x_356); lean_inc(x_705); x_760 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_760, 0, x_705); -lean_ctor_set(x_760, 1, x_338); -lean_inc(x_337); +lean_ctor_set(x_760, 1, x_360); +lean_inc(x_338); lean_inc(x_705); x_761 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_761, 0, x_705); -lean_ctor_set(x_761, 1, x_337); -lean_ctor_set(x_761, 2, x_370); +lean_ctor_set(x_761, 1, x_338); +lean_ctor_set(x_761, 2, x_362); x_762 = 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__2_spec__2___closed__34; x_763 = 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__2_spec__2___closed__35; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_764 = l_Lean_addMacroScope(x_346, x_763, x_369); lean_inc(x_356); -x_764 = l_Lean_addMacroScope(x_356, x_763, x_341); -lean_inc(x_352); lean_inc(x_705); x_765 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_765, 0, x_705); lean_ctor_set(x_765, 1, x_762); lean_ctor_set(x_765, 2, x_764); -lean_ctor_set(x_765, 3, x_352); +lean_ctor_set(x_765, 3, x_356); lean_inc_ref(x_761); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_705); -x_766 = l_Lean_Syntax_node2(x_705, x_351, x_765, x_761); +x_766 = l_Lean_Syntax_node2(x_705, x_347, x_765, x_761); lean_inc(x_705); x_767 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_767, 0, x_705); -lean_ctor_set(x_767, 1, x_372); +lean_ctor_set(x_767, 1, x_345); lean_inc_ref(x_761); lean_inc_ref(x_767); -lean_inc(x_342); +lean_inc(x_357); lean_inc(x_705); -x_768 = l_Lean_Syntax_node3(x_705, x_342, x_767, x_761, x_361); +x_768 = l_Lean_Syntax_node3(x_705, x_357, x_767, x_761, x_340); lean_inc_ref_n(x_761, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_705); -x_769 = l_Lean_Syntax_node3(x_705, x_337, x_761, x_761, x_768); +x_769 = l_Lean_Syntax_node3(x_705, x_338, x_761, x_761, x_768); lean_inc(x_769); -lean_inc(x_344); +lean_inc(x_349); lean_inc(x_705); -x_770 = l_Lean_Syntax_node2(x_705, x_344, x_766, x_769); +x_770 = l_Lean_Syntax_node2(x_705, x_349, x_766, x_769); x_771 = 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__2_spec__2___closed__38; x_772 = 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__2_spec__2___closed__39; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_773 = l_Lean_addMacroScope(x_346, x_772, x_369); lean_inc(x_356); -x_773 = l_Lean_addMacroScope(x_356, x_772, x_341); -lean_inc(x_352); lean_inc(x_705); x_774 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_774, 0, x_705); lean_ctor_set(x_774, 1, x_771); lean_ctor_set(x_774, 2, x_773); -lean_ctor_set(x_774, 3, x_352); +lean_ctor_set(x_774, 3, x_356); lean_inc_ref(x_761); -lean_inc(x_351); +lean_inc(x_347); lean_inc(x_705); -x_775 = l_Lean_Syntax_node2(x_705, x_351, x_774, x_761); -lean_inc(x_344); +x_775 = l_Lean_Syntax_node2(x_705, x_347, x_774, x_761); +lean_inc(x_349); lean_inc(x_705); -x_776 = l_Lean_Syntax_node2(x_705, x_344, x_775, x_769); +x_776 = l_Lean_Syntax_node2(x_705, x_349, x_775, x_769); x_777 = 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__4_spec__4___closed__21; x_778 = 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__4_spec__4___closed__22; -lean_inc(x_341); +lean_inc(x_369); +lean_inc(x_346); +x_779 = l_Lean_addMacroScope(x_346, x_778, x_369); lean_inc(x_356); -x_779 = l_Lean_addMacroScope(x_356, x_778, x_341); -lean_inc(x_352); lean_inc(x_705); x_780 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_780, 0, x_705); lean_ctor_set(x_780, 1, x_777); lean_ctor_set(x_780, 2, x_779); -lean_ctor_set(x_780, 3, x_352); +lean_ctor_set(x_780, 3, x_356); lean_inc_ref(x_761); lean_inc(x_705); -x_781 = l_Lean_Syntax_node2(x_705, x_351, x_780, x_761); +x_781 = l_Lean_Syntax_node2(x_705, x_347, x_780, x_761); x_782 = 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__4_spec__4___closed__23; x_783 = 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__4_spec__4___closed__24; -lean_inc(x_341); -lean_inc(x_356); -x_784 = l_Lean_addMacroScope(x_356, x_783, x_341); +lean_inc(x_369); +lean_inc(x_346); +x_784 = l_Lean_addMacroScope(x_346, x_783, x_369); x_785 = 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__2_spec__2___closed__48; -lean_inc(x_371); +lean_inc(x_339); x_786 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_786, 0, x_785); -lean_ctor_set(x_786, 1, x_371); -lean_inc(x_352); +lean_ctor_set(x_786, 1, x_339); +lean_inc(x_356); x_787 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_787, 0, x_786); -lean_ctor_set(x_787, 1, x_352); +lean_ctor_set(x_787, 1, x_356); lean_inc(x_705); x_788 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_788, 0, x_705); @@ -7025,40 +7025,40 @@ lean_ctor_set(x_788, 2, x_784); lean_ctor_set(x_788, 3, x_787); lean_inc_ref(x_761); lean_inc(x_705); -x_789 = l_Lean_Syntax_node3(x_705, x_342, x_767, x_761, x_788); +x_789 = l_Lean_Syntax_node3(x_705, x_357, x_767, x_761, x_788); lean_inc_ref_n(x_761, 2); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_705); -x_790 = l_Lean_Syntax_node3(x_705, x_337, x_761, x_761, x_789); +x_790 = l_Lean_Syntax_node3(x_705, x_338, x_761, x_761, x_789); lean_inc(x_705); -x_791 = l_Lean_Syntax_node2(x_705, x_344, x_781, x_790); +x_791 = l_Lean_Syntax_node2(x_705, x_349, x_781, x_790); lean_inc_ref_n(x_761, 3); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_705); -x_792 = l_Lean_Syntax_node6(x_705, x_337, x_770, x_761, x_776, x_761, x_791, x_761); +x_792 = l_Lean_Syntax_node6(x_705, x_338, x_770, x_761, x_776, x_761, x_791, x_761); lean_inc(x_705); -x_793 = l_Lean_Syntax_node1(x_705, x_339, x_792); +x_793 = l_Lean_Syntax_node1(x_705, x_363, x_792); lean_inc_ref(x_761); lean_inc(x_705); -x_794 = l_Lean_Syntax_node1(x_705, x_358, x_761); +x_794 = l_Lean_Syntax_node1(x_705, x_355, x_761); lean_inc(x_705); x_795 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_795, 0, x_705); -lean_ctor_set(x_795, 1, x_353); +lean_ctor_set(x_795, 1, x_361); x_796 = 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__2_spec__2___closed__1; x_797 = 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__2_spec__2___closed__2; x_798 = 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__2_spec__2___closed__3; -x_799 = l_Lean_addMacroScope(x_356, x_798, x_341); -x_800 = l_Lean_Name_mkStr2(x_355, x_796); +x_799 = l_Lean_addMacroScope(x_346, x_798, x_369); +x_800 = l_Lean_Name_mkStr2(x_341, x_796); lean_inc(x_800); x_801 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_801, 0, x_800); -lean_ctor_set(x_801, 1, x_371); +lean_ctor_set(x_801, 1, x_339); x_802 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_802, 0, x_800); x_803 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_803, 0, x_802); -lean_ctor_set(x_803, 1, x_352); +lean_ctor_set(x_803, 1, x_356); x_804 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_804, 0, x_801); lean_ctor_set(x_804, 1, x_803); @@ -7068,17 +7068,17 @@ lean_ctor_set(x_805, 0, x_705); lean_ctor_set(x_805, 1, x_797); lean_ctor_set(x_805, 2, x_799); lean_ctor_set(x_805, 3, x_804); -lean_inc(x_337); +lean_inc(x_338); lean_inc(x_705); -x_806 = l_Lean_Syntax_node2(x_705, x_337, x_795, x_805); +x_806 = l_Lean_Syntax_node2(x_705, x_338, x_795, x_805); lean_inc(x_705); x_807 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_807, 0, x_705); -lean_ctor_set(x_807, 1, x_354); +lean_ctor_set(x_807, 1, x_342); lean_inc(x_705); -x_808 = l_Lean_Syntax_node6(x_705, x_366, x_760, x_761, x_793, x_794, x_806, x_807); +x_808 = l_Lean_Syntax_node6(x_705, x_351, x_760, x_761, x_793, x_794, x_806, x_807); lean_inc(x_705); -x_809 = l_Lean_Syntax_node1(x_705, x_337, x_808); +x_809 = l_Lean_Syntax_node1(x_705, x_338, x_808); x_810 = l_Lean_Syntax_node4(x_705, x_650, x_701, x_755, x_759, x_809); x_811 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_811, 0, x_754); @@ -7091,7 +7091,7 @@ goto block_16; block_855: { lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; -x_847 = 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__2_spec__2___lam__1(x_17, x_823, x_9, x_10); +x_847 = 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__2_spec__2___lam__1(x_17, x_820, x_9, x_10); x_848 = lean_ctor_get(x_847, 0); lean_inc(x_848); x_849 = lean_ctor_get(x_847, 1); @@ -7099,52 +7099,52 @@ lean_inc(x_849); lean_dec_ref(x_847); x_850 = l_Lean_mkIdentFrom(x_22, x_846, x_17); lean_dec(x_22); -lean_inc_ref(x_843); -lean_inc(x_814); +lean_inc_ref(x_837); +lean_inc(x_813); lean_inc(x_848); x_851 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_851, 0, x_848); -lean_ctor_set(x_851, 1, x_814); -lean_ctor_set(x_851, 2, x_843); +lean_ctor_set(x_851, 1, x_813); +lean_ctor_set(x_851, 2, x_837); if (lean_obj_tag(x_3) == 0) { lean_object* x_852; x_852 = 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__2_spec__2___closed__53; -x_336 = x_813; -x_337 = x_814; -x_338 = x_815; -x_339 = x_818; -x_340 = x_817; +x_336 = x_848; +x_337 = x_850; +x_338 = x_813; +x_339 = x_814; +x_340 = x_815; x_341 = x_816; -x_342 = x_819; -x_343 = x_820; -x_344 = x_822; -x_345 = x_821; -x_346 = x_848; -x_347 = x_823; -x_348 = x_824; -x_349 = x_825; -x_350 = x_826; -x_351 = x_827; -x_352 = x_828; -x_353 = x_829; -x_354 = x_830; -x_355 = x_831; -x_356 = x_832; -x_357 = x_850; +x_342 = x_817; +x_343 = x_818; +x_344 = x_820; +x_345 = x_819; +x_346 = x_821; +x_347 = x_822; +x_348 = x_823; +x_349 = x_824; +x_350 = x_825; +x_351 = x_826; +x_352 = x_827; +x_353 = x_828; +x_354 = x_829; +x_355 = x_830; +x_356 = x_831; +x_357 = x_832; x_358 = x_833; -x_359 = x_849; -x_360 = x_834; -x_361 = x_835; +x_359 = x_834; +x_360 = x_835; +x_361 = x_836; x_362 = x_837; -x_363 = x_836; -x_364 = x_838; -x_365 = x_839; -x_366 = x_840; -x_367 = x_851; -x_368 = x_841; +x_363 = x_840; +x_364 = x_839; +x_365 = x_838; +x_366 = x_841; +x_367 = x_843; +x_368 = x_849; x_369 = x_842; -x_370 = x_843; +x_370 = x_851; x_371 = x_845; x_372 = x_844; x_373 = x_852; @@ -7156,41 +7156,41 @@ lean_object* x_853; lean_object* x_854; x_853 = lean_ctor_get(x_3, 0); lean_inc(x_853); x_854 = l_Array_mkArray1___redArg(x_853); -x_336 = x_813; -x_337 = x_814; -x_338 = x_815; -x_339 = x_818; -x_340 = x_817; +x_336 = x_848; +x_337 = x_850; +x_338 = x_813; +x_339 = x_814; +x_340 = x_815; x_341 = x_816; -x_342 = x_819; -x_343 = x_820; -x_344 = x_822; -x_345 = x_821; -x_346 = x_848; -x_347 = x_823; -x_348 = x_824; -x_349 = x_825; -x_350 = x_826; -x_351 = x_827; -x_352 = x_828; -x_353 = x_829; -x_354 = x_830; -x_355 = x_831; -x_356 = x_832; -x_357 = x_850; +x_342 = x_817; +x_343 = x_818; +x_344 = x_820; +x_345 = x_819; +x_346 = x_821; +x_347 = x_822; +x_348 = x_823; +x_349 = x_824; +x_350 = x_825; +x_351 = x_826; +x_352 = x_827; +x_353 = x_828; +x_354 = x_829; +x_355 = x_830; +x_356 = x_831; +x_357 = x_832; x_358 = x_833; -x_359 = x_849; -x_360 = x_834; -x_361 = x_835; +x_359 = x_834; +x_360 = x_835; +x_361 = x_836; x_362 = x_837; -x_363 = x_836; -x_364 = x_838; -x_365 = x_839; -x_366 = x_840; -x_367 = x_851; -x_368 = x_841; +x_363 = x_840; +x_364 = x_839; +x_365 = x_838; +x_366 = x_841; +x_367 = x_843; +x_368 = x_849; x_369 = x_842; -x_370 = x_843; +x_370 = x_851; x_371 = x_845; x_372 = x_844; x_373 = x_854; @@ -7200,428 +7200,428 @@ goto block_812; block_1049: { lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; -lean_inc_ref(x_871); -x_874 = l_Array_append___redArg(x_871, x_873); +lean_inc_ref(x_861); +x_874 = l_Array_append___redArg(x_861, x_873); lean_dec_ref(x_873); lean_inc(x_857); -lean_inc(x_872); +lean_inc(x_860); x_875 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_875, 0, x_872); +lean_ctor_set(x_875, 0, x_860); lean_ctor_set(x_875, 1, x_857); lean_ctor_set(x_875, 2, x_874); -lean_inc_n(x_868, 6); +lean_inc_n(x_859, 6); +lean_inc(x_862); lean_inc(x_860); -lean_inc(x_872); -x_876 = l_Lean_Syntax_node7(x_872, x_860, x_868, x_868, x_875, x_868, x_868, x_868, x_868); +x_876 = l_Lean_Syntax_node7(x_860, x_862, x_859, x_859, x_875, x_859, x_859, x_859, x_859); x_877 = 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__4_spec__4___closed__25; -lean_inc_ref(x_861); +lean_inc_ref(x_867); +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_878 = l_Lean_Name_mkStr4(x_869, x_863, x_861, x_877); +x_878 = l_Lean_Name_mkStr4(x_863, x_864, x_867, x_877); x_879 = 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__4_spec__4___closed__26; -lean_inc(x_872); +lean_inc(x_860); x_880 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_880, 0, x_872); +lean_ctor_set(x_880, 0, x_860); lean_ctor_set(x_880, 1, x_879); x_881 = l_Lake_configDecl___closed__8; -lean_inc_ref(x_861); +lean_inc_ref(x_867); +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_882 = l_Lean_Name_mkStr4(x_869, x_863, x_861, x_881); -lean_inc(x_868); -lean_inc(x_862); +x_882 = l_Lean_Name_mkStr4(x_863, x_864, x_867, x_881); +lean_inc(x_859); +lean_inc(x_871); lean_inc(x_882); -lean_inc(x_872); -x_883 = l_Lean_Syntax_node2(x_872, x_882, x_862, x_868); +lean_inc(x_860); +x_883 = l_Lean_Syntax_node2(x_860, x_882, x_871, x_859); x_884 = 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__4_spec__4___closed__27; -lean_inc_ref(x_861); +lean_inc_ref(x_867); +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_885 = l_Lean_Name_mkStr4(x_869, x_863, x_861, x_884); +x_885 = l_Lean_Name_mkStr4(x_863, x_864, x_867, x_884); x_886 = l_Lake_configDecl___closed__26; x_887 = 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__2_spec__2___closed__7; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_888 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_887); +x_888 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_887); x_889 = 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__2_spec__2___closed__8; -lean_inc(x_872); +lean_inc(x_860); x_890 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_890, 0, x_872); +lean_ctor_set(x_890, 0, x_860); lean_ctor_set(x_890, 1, x_889); x_891 = 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__2_spec__2___closed__9; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_892 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_891); +x_892 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_891); x_893 = 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__4_spec__4___closed__29; x_894 = 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__4_spec__4___closed__30; -lean_inc(x_858); -lean_inc(x_859); -x_895 = l_Lean_addMacroScope(x_859, x_894, x_858); +lean_inc(x_868); +lean_inc(x_866); +x_895 = l_Lean_addMacroScope(x_866, x_894, x_868); x_896 = l_Lake_configField___closed__1; x_897 = lean_box(0); x_898 = 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__4_spec__4___closed__35; -lean_inc(x_872); +lean_inc(x_860); x_899 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_899, 0, x_872); +lean_ctor_set(x_899, 0, x_860); lean_ctor_set(x_899, 1, x_893); lean_ctor_set(x_899, 2, x_895); lean_ctor_set(x_899, 3, x_898); lean_inc(x_24); lean_inc(x_1); lean_inc(x_857); -lean_inc(x_872); -x_900 = l_Lean_Syntax_node2(x_872, x_857, x_1, x_24); +lean_inc(x_860); +x_900 = l_Lean_Syntax_node2(x_860, x_857, x_1, x_24); lean_inc(x_892); -lean_inc(x_872); -x_901 = l_Lean_Syntax_node2(x_872, x_892, x_899, x_900); +lean_inc(x_860); +x_901 = l_Lean_Syntax_node2(x_860, x_892, x_899, x_900); lean_inc(x_888); -lean_inc(x_872); -x_902 = l_Lean_Syntax_node2(x_872, x_888, x_890, x_901); +lean_inc(x_860); +x_902 = l_Lean_Syntax_node2(x_860, x_888, x_890, x_901); lean_inc(x_857); -lean_inc(x_872); -x_903 = l_Lean_Syntax_node1(x_872, x_857, x_902); -lean_inc(x_868); -lean_inc(x_872); -x_904 = l_Lean_Syntax_node2(x_872, x_885, x_868, x_903); +lean_inc(x_860); +x_903 = l_Lean_Syntax_node1(x_860, x_857, x_902); +lean_inc(x_859); +lean_inc(x_860); +x_904 = l_Lean_Syntax_node2(x_860, x_885, x_859, x_903); x_905 = 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__4_spec__4___closed__36; -lean_inc_ref(x_861); +lean_inc_ref(x_867); +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_906 = l_Lean_Name_mkStr4(x_869, x_863, x_861, x_905); +x_906 = l_Lean_Name_mkStr4(x_863, x_864, x_867, x_905); x_907 = 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__4_spec__4___closed__37; -lean_inc(x_872); +lean_inc(x_860); x_908 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_908, 0, x_872); +lean_ctor_set(x_908, 0, x_860); lean_ctor_set(x_908, 1, x_907); x_909 = 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__2_spec__2___closed__32; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_910 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_909); +x_910 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_909); x_911 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_912 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_911); +x_912 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_911); x_913 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_914 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_913); +x_914 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_913); x_915 = 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__4_spec__4___closed__39; x_916 = 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__4_spec__4___closed__40; -lean_inc(x_858); -lean_inc(x_859); -x_917 = l_Lean_addMacroScope(x_859, x_916, x_858); +lean_inc(x_868); +lean_inc(x_866); +x_917 = l_Lean_addMacroScope(x_866, x_916, x_868); x_918 = 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__4_spec__4___closed__44; -lean_inc(x_872); +lean_inc(x_860); x_919 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_919, 0, x_872); +lean_ctor_set(x_919, 0, x_860); lean_ctor_set(x_919, 1, x_915); lean_ctor_set(x_919, 2, x_917); lean_ctor_set(x_919, 3, x_918); -lean_inc(x_868); +lean_inc(x_859); lean_inc(x_914); -lean_inc(x_872); -x_920 = l_Lean_Syntax_node2(x_872, x_914, x_919, x_868); +lean_inc(x_860); +x_920 = l_Lean_Syntax_node2(x_860, x_914, x_919, x_859); x_921 = 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__4_spec__4___closed__46; x_922 = 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__4_spec__4___closed__47; -lean_inc(x_858); -lean_inc(x_859); -x_923 = l_Lean_addMacroScope(x_859, x_922, x_858); -lean_inc(x_872); +lean_inc(x_868); +lean_inc(x_866); +x_923 = l_Lean_addMacroScope(x_866, x_922, x_868); +lean_inc(x_860); x_924 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_924, 0, x_872); +lean_ctor_set(x_924, 0, x_860); lean_ctor_set(x_924, 1, x_921); lean_ctor_set(x_924, 2, x_923); lean_ctor_set(x_924, 3, x_897); lean_inc_ref(x_924); lean_inc(x_857); -lean_inc(x_872); -x_925 = l_Lean_Syntax_node1(x_872, x_857, x_924); +lean_inc(x_860); +x_925 = l_Lean_Syntax_node1(x_860, x_857, x_924); x_926 = 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__2_spec__2___closed__36; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_927 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_926); +x_927 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_926); x_928 = 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__2_spec__2___closed__19; -lean_inc(x_872); +lean_inc(x_860); x_929 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_929, 0, x_872); +lean_ctor_set(x_929, 0, x_860); lean_ctor_set(x_929, 1, x_928); x_930 = 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__4_spec__4___closed__48; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_931 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_930); +x_931 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_930); x_932 = 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__4_spec__4___closed__49; -lean_inc(x_872); +lean_inc(x_860); x_933 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_933, 0, x_872); +lean_ctor_set(x_933, 0, x_860); lean_ctor_set(x_933, 1, x_932); lean_inc(x_22); lean_inc_ref(x_924); -lean_inc(x_872); -x_934 = l_Lean_Syntax_node3(x_872, x_931, x_924, x_933, x_22); +lean_inc(x_860); +x_934 = l_Lean_Syntax_node3(x_860, x_931, x_924, x_933, x_22); lean_inc(x_934); -lean_inc(x_868); +lean_inc(x_859); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_872); -x_935 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_934); -lean_inc(x_868); +lean_inc(x_860); +x_935 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_934); +lean_inc(x_859); lean_inc(x_925); lean_inc(x_857); -lean_inc(x_872); -x_936 = l_Lean_Syntax_node3(x_872, x_857, x_925, x_868, x_935); +lean_inc(x_860); +x_936 = l_Lean_Syntax_node3(x_860, x_857, x_925, x_859, x_935); lean_inc(x_912); -lean_inc(x_872); -x_937 = l_Lean_Syntax_node2(x_872, x_912, x_920, x_936); +lean_inc(x_860); +x_937 = l_Lean_Syntax_node2(x_860, x_912, x_920, x_936); x_938 = 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__4_spec__4___closed__51; x_939 = 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__4_spec__4___closed__52; -lean_inc(x_858); -lean_inc(x_859); -x_940 = l_Lean_addMacroScope(x_859, x_939, x_858); +lean_inc(x_868); +lean_inc(x_866); +x_940 = l_Lean_addMacroScope(x_866, x_939, x_868); x_941 = 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__4_spec__4___closed__56; -lean_inc(x_872); +lean_inc(x_860); x_942 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_942, 0, x_872); +lean_ctor_set(x_942, 0, x_860); lean_ctor_set(x_942, 1, x_938); lean_ctor_set(x_942, 2, x_940); lean_ctor_set(x_942, 3, x_941); -lean_inc(x_868); +lean_inc(x_859); lean_inc(x_914); -lean_inc(x_872); -x_943 = l_Lean_Syntax_node2(x_872, x_914, x_942, x_868); +lean_inc(x_860); +x_943 = l_Lean_Syntax_node2(x_860, x_914, x_942, x_859); x_944 = 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__4_spec__4___closed__58; x_945 = 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__4_spec__4___closed__59; -lean_inc(x_858); -lean_inc(x_859); -x_946 = l_Lean_addMacroScope(x_859, x_945, x_858); -lean_inc(x_872); +lean_inc(x_868); +lean_inc(x_866); +x_946 = l_Lean_addMacroScope(x_866, x_945, x_868); +lean_inc(x_860); x_947 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_947, 0, x_872); +lean_ctor_set(x_947, 0, x_860); lean_ctor_set(x_947, 1, x_944); lean_ctor_set(x_947, 2, x_946); lean_ctor_set(x_947, 3, x_897); lean_inc_ref(x_924); lean_inc_ref(x_947); lean_inc(x_857); -lean_inc(x_872); -x_948 = l_Lean_Syntax_node2(x_872, x_857, x_947, x_924); +lean_inc(x_860); +x_948 = l_Lean_Syntax_node2(x_860, x_857, x_947, x_924); x_949 = 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__2_spec__2___closed__30; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_950 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_949); +x_950 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_949); x_951 = 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__2_spec__2___closed__31; -lean_inc(x_872); +lean_inc(x_860); x_952 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_952, 0, x_872); +lean_ctor_set(x_952, 0, x_860); lean_ctor_set(x_952, 1, x_951); x_953 = 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__4_spec__4___closed__60; -lean_inc(x_872); +lean_inc(x_860); x_954 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_954, 0, x_872); +lean_ctor_set(x_954, 0, x_860); lean_ctor_set(x_954, 1, x_953); lean_inc(x_857); -lean_inc(x_872); -x_955 = l_Lean_Syntax_node2(x_872, x_857, x_925, x_954); +lean_inc(x_860); +x_955 = l_Lean_Syntax_node2(x_860, x_857, x_925, x_954); x_956 = lean_box(0); x_957 = l_Lean_SourceInfo_fromRef(x_956, x_17); -lean_inc_ref(x_871); +lean_inc_ref(x_861); lean_inc(x_857); lean_inc(x_957); x_958 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_958, 0, x_957); lean_ctor_set(x_958, 1, x_857); -lean_ctor_set(x_958, 2, x_871); +lean_ctor_set(x_958, 2, x_861); lean_inc(x_22); lean_inc(x_914); x_959 = l_Lean_Syntax_node2(x_957, x_914, x_22, x_958); -lean_inc(x_868); +lean_inc(x_859); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_872); -x_960 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_947); -lean_inc_n(x_868, 2); +lean_inc(x_860); +x_960 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_947); +lean_inc_n(x_859, 2); lean_inc(x_857); -lean_inc(x_872); -x_961 = l_Lean_Syntax_node3(x_872, x_857, x_868, x_868, x_960); +lean_inc(x_860); +x_961 = l_Lean_Syntax_node3(x_860, x_857, x_859, x_859, x_960); lean_inc(x_959); lean_inc(x_912); -lean_inc(x_872); -x_962 = l_Lean_Syntax_node2(x_872, x_912, x_959, x_961); +lean_inc(x_860); +x_962 = l_Lean_Syntax_node2(x_860, x_912, x_959, x_961); lean_inc(x_857); -lean_inc(x_872); -x_963 = l_Lean_Syntax_node1(x_872, x_857, x_962); +lean_inc(x_860); +x_963 = l_Lean_Syntax_node1(x_860, x_857, x_962); lean_inc(x_910); -lean_inc(x_872); -x_964 = l_Lean_Syntax_node1(x_872, x_910, x_963); +lean_inc(x_860); +x_964 = l_Lean_Syntax_node1(x_860, x_910, x_963); x_965 = 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__2_spec__2___closed__0; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_966 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_965); -lean_inc(x_868); +x_966 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_965); +lean_inc(x_859); lean_inc(x_966); -lean_inc(x_872); -x_967 = l_Lean_Syntax_node1(x_872, x_966, x_868); +lean_inc(x_860); +x_967 = l_Lean_Syntax_node1(x_860, x_966, x_859); x_968 = 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__2_spec__2___closed__4; -lean_inc(x_872); +lean_inc(x_860); x_969 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_969, 0, x_872); +lean_ctor_set(x_969, 0, x_860); lean_ctor_set(x_969, 1, x_968); lean_inc_ref(x_969); -lean_inc(x_868); +lean_inc(x_859); lean_inc(x_967); lean_inc(x_955); lean_inc_ref(x_952); lean_inc(x_950); -lean_inc(x_872); -x_970 = l_Lean_Syntax_node6(x_872, x_950, x_952, x_955, x_964, x_967, x_868, x_969); -lean_inc(x_868); +lean_inc(x_860); +x_970 = l_Lean_Syntax_node6(x_860, x_950, x_952, x_955, x_964, x_967, x_859, x_969); +lean_inc(x_859); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_872); -x_971 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_970); -lean_inc(x_868); +lean_inc(x_860); +x_971 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_970); +lean_inc(x_859); lean_inc(x_857); -lean_inc(x_872); -x_972 = l_Lean_Syntax_node3(x_872, x_857, x_948, x_868, x_971); +lean_inc(x_860); +x_972 = l_Lean_Syntax_node3(x_860, x_857, x_948, x_859, x_971); lean_inc(x_912); -lean_inc(x_872); -x_973 = l_Lean_Syntax_node2(x_872, x_912, x_943, x_972); +lean_inc(x_860); +x_973 = l_Lean_Syntax_node2(x_860, x_912, x_943, x_972); x_974 = 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__4_spec__4___closed__62; x_975 = 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__4_spec__4___closed__63; -lean_inc(x_858); -lean_inc(x_859); -x_976 = l_Lean_addMacroScope(x_859, x_975, x_858); +lean_inc(x_868); +lean_inc(x_866); +x_976 = l_Lean_addMacroScope(x_866, x_975, x_868); x_977 = 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__4_spec__4___closed__65; -lean_inc(x_872); +lean_inc(x_860); x_978 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_978, 0, x_872); +lean_ctor_set(x_978, 0, x_860); lean_ctor_set(x_978, 1, x_974); lean_ctor_set(x_978, 2, x_976); lean_ctor_set(x_978, 3, x_977); -lean_inc(x_868); +lean_inc(x_859); lean_inc(x_914); -lean_inc(x_872); -x_979 = l_Lean_Syntax_node2(x_872, x_914, x_978, x_868); +lean_inc(x_860); +x_979 = l_Lean_Syntax_node2(x_860, x_914, x_978, x_859); x_980 = 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__4_spec__4___closed__67; x_981 = 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__4_spec__4___closed__68; -lean_inc(x_858); -lean_inc(x_859); -x_982 = l_Lean_addMacroScope(x_859, x_981, x_858); -lean_inc(x_872); +lean_inc(x_868); +lean_inc(x_866); +x_982 = l_Lean_addMacroScope(x_866, x_981, x_868); +lean_inc(x_860); x_983 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_983, 0, x_872); +lean_ctor_set(x_983, 0, x_860); lean_ctor_set(x_983, 1, x_980); lean_ctor_set(x_983, 2, x_982); lean_ctor_set(x_983, 3, x_897); lean_inc_ref(x_983); lean_inc(x_857); -lean_inc(x_872); -x_984 = l_Lean_Syntax_node2(x_872, x_857, x_983, x_924); +lean_inc(x_860); +x_984 = l_Lean_Syntax_node2(x_860, x_857, x_983, x_924); lean_inc(x_857); -lean_inc(x_872); -x_985 = l_Lean_Syntax_node1(x_872, x_857, x_934); +lean_inc(x_860); +x_985 = l_Lean_Syntax_node1(x_860, x_857, x_934); lean_inc(x_892); -lean_inc(x_872); -x_986 = l_Lean_Syntax_node2(x_872, x_892, x_983, x_985); -lean_inc(x_868); +lean_inc(x_860); +x_986 = l_Lean_Syntax_node2(x_860, x_892, x_983, x_985); +lean_inc(x_859); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_872); -x_987 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_986); -lean_inc_n(x_868, 2); +lean_inc(x_860); +x_987 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_986); +lean_inc_n(x_859, 2); lean_inc(x_857); -lean_inc(x_872); -x_988 = l_Lean_Syntax_node3(x_872, x_857, x_868, x_868, x_987); +lean_inc(x_860); +x_988 = l_Lean_Syntax_node3(x_860, x_857, x_859, x_859, x_987); lean_inc(x_912); -lean_inc(x_872); -x_989 = l_Lean_Syntax_node2(x_872, x_912, x_959, x_988); +lean_inc(x_860); +x_989 = l_Lean_Syntax_node2(x_860, x_912, x_959, x_988); lean_inc(x_857); -lean_inc(x_872); -x_990 = l_Lean_Syntax_node1(x_872, x_857, x_989); +lean_inc(x_860); +x_990 = l_Lean_Syntax_node1(x_860, x_857, x_989); lean_inc(x_910); -lean_inc(x_872); -x_991 = l_Lean_Syntax_node1(x_872, x_910, x_990); -lean_inc(x_868); +lean_inc(x_860); +x_991 = l_Lean_Syntax_node1(x_860, x_910, x_990); +lean_inc(x_859); lean_inc(x_950); -lean_inc(x_872); -x_992 = l_Lean_Syntax_node6(x_872, x_950, x_952, x_955, x_991, x_967, x_868, x_969); -lean_inc(x_868); +lean_inc(x_860); +x_992 = l_Lean_Syntax_node6(x_860, x_950, x_952, x_955, x_991, x_967, x_859, x_969); +lean_inc(x_859); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_872); -x_993 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_992); -lean_inc(x_868); +lean_inc(x_860); +x_993 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_992); +lean_inc(x_859); lean_inc(x_857); -lean_inc(x_872); -x_994 = l_Lean_Syntax_node3(x_872, x_857, x_984, x_868, x_993); +lean_inc(x_860); +x_994 = l_Lean_Syntax_node3(x_860, x_857, x_984, x_859, x_993); lean_inc(x_912); -lean_inc(x_872); -x_995 = l_Lean_Syntax_node2(x_872, x_912, x_979, x_994); +lean_inc(x_860); +x_995 = l_Lean_Syntax_node2(x_860, x_912, x_979, x_994); x_996 = 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__4_spec__4___closed__70; x_997 = 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__4_spec__4___closed__71; -lean_inc(x_858); -lean_inc(x_859); -x_998 = l_Lean_addMacroScope(x_859, x_997, x_858); -lean_inc(x_872); +lean_inc(x_868); +lean_inc(x_866); +x_998 = l_Lean_addMacroScope(x_866, x_997, x_868); +lean_inc(x_860); x_999 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_999, 0, x_872); +lean_ctor_set(x_999, 0, x_860); lean_ctor_set(x_999, 1, x_996); lean_ctor_set(x_999, 2, x_998); lean_ctor_set(x_999, 3, x_897); -lean_inc(x_868); +lean_inc(x_859); lean_inc(x_914); -lean_inc(x_872); -x_1000 = l_Lean_Syntax_node2(x_872, x_914, x_999, x_868); +lean_inc(x_860); +x_1000 = l_Lean_Syntax_node2(x_860, x_914, x_999, x_859); x_1001 = 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__4_spec__4___closed__72; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_1002 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_1001); -lean_inc(x_872); +x_1002 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_1001); +lean_inc(x_860); x_1003 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1003, 0, x_872); +lean_ctor_set(x_1003, 0, x_860); lean_ctor_set(x_1003, 1, x_1001); x_1004 = 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__4_spec__4___closed__73; +lean_inc_ref(x_864); lean_inc_ref(x_863); -lean_inc_ref(x_869); -x_1005 = l_Lean_Name_mkStr4(x_869, x_863, x_886, x_1004); +x_1005 = l_Lean_Name_mkStr4(x_863, x_864, x_886, x_1004); lean_inc(x_2); lean_inc(x_857); -lean_inc(x_872); -x_1006 = l_Lean_Syntax_node1(x_872, x_857, x_2); +lean_inc(x_860); +x_1006 = l_Lean_Syntax_node1(x_860, x_857, x_2); x_1007 = 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__4_spec__4___closed__74; -lean_inc(x_872); +lean_inc(x_860); x_1008 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1008, 0, x_872); +lean_ctor_set(x_1008, 0, x_860); lean_ctor_set(x_1008, 1, x_1007); -lean_inc(x_868); -lean_inc(x_872); -x_1009 = l_Lean_Syntax_node4(x_872, x_1005, x_1006, x_868, x_1008, x_25); -lean_inc(x_872); -x_1010 = l_Lean_Syntax_node2(x_872, x_1002, x_1003, x_1009); -lean_inc(x_868); +lean_inc(x_859); +lean_inc(x_860); +x_1009 = l_Lean_Syntax_node4(x_860, x_1005, x_1006, x_859, x_1008, x_25); +lean_inc(x_860); +x_1010 = l_Lean_Syntax_node2(x_860, x_1002, x_1003, x_1009); +lean_inc(x_859); lean_inc(x_927); -lean_inc(x_872); -x_1011 = l_Lean_Syntax_node3(x_872, x_927, x_929, x_868, x_1010); -lean_inc_n(x_868, 2); +lean_inc(x_860); +x_1011 = l_Lean_Syntax_node3(x_860, x_927, x_929, x_859, x_1010); +lean_inc_n(x_859, 2); lean_inc(x_857); -lean_inc(x_872); -x_1012 = l_Lean_Syntax_node3(x_872, x_857, x_868, x_868, x_1011); +lean_inc(x_860); +x_1012 = l_Lean_Syntax_node3(x_860, x_857, x_859, x_859, x_1011); lean_inc(x_912); -lean_inc(x_872); -x_1013 = l_Lean_Syntax_node2(x_872, x_912, x_1000, x_1012); -lean_inc_n(x_868, 3); +lean_inc(x_860); +x_1013 = l_Lean_Syntax_node2(x_860, x_912, x_1000, x_1012); +lean_inc_n(x_859, 3); lean_inc(x_857); -lean_inc(x_872); -x_1014 = l_Lean_Syntax_node7(x_872, x_857, x_937, x_868, x_973, x_868, x_995, x_868, x_1013); +lean_inc(x_860); +x_1014 = l_Lean_Syntax_node7(x_860, x_857, x_937, x_859, x_973, x_859, x_995, x_859, x_1013); lean_inc(x_910); -lean_inc(x_872); -x_1015 = l_Lean_Syntax_node1(x_872, x_910, x_1014); -lean_inc(x_868); -lean_inc(x_872); -x_1016 = l_Lean_Syntax_node3(x_872, x_906, x_908, x_1015, x_868); -lean_inc(x_872); -x_1017 = l_Lean_Syntax_node5(x_872, x_878, x_880, x_883, x_904, x_1016, x_868); -lean_inc(x_867); -x_1018 = l_Lean_Syntax_node2(x_872, x_867, x_876, x_1017); +lean_inc(x_860); +x_1015 = l_Lean_Syntax_node1(x_860, x_910, x_1014); +lean_inc(x_859); +lean_inc(x_860); +x_1016 = l_Lean_Syntax_node3(x_860, x_906, x_908, x_1015, x_859); +lean_inc(x_860); +x_1017 = l_Lean_Syntax_node5(x_860, x_878, x_880, x_883, x_904, x_1016, x_859); +lean_inc(x_858); +x_1018 = l_Lean_Syntax_node2(x_860, x_858, x_876, x_1017); x_1019 = lean_array_push(x_18, x_1018); lean_inc(x_856); lean_inc(x_22); @@ -7634,39 +7634,39 @@ if (x_1021 == 0) { lean_object* x_1022; x_1022 = 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__2_spec__2___lam__0(x_4, x_856); -x_293 = x_1019; -x_294 = x_857; -x_295 = x_951; -x_296 = x_886; -x_297 = x_858; -x_298 = x_910; -x_299 = x_927; -x_300 = x_882; -x_301 = x_862; -x_302 = x_912; -x_303 = x_865; -x_304 = x_867; -x_305 = x_888; -x_306 = x_869; -x_307 = x_914; -x_308 = x_897; -x_309 = x_889; -x_310 = x_968; -x_311 = x_896; -x_312 = x_859; -x_313 = x_966; -x_314 = x_860; -x_315 = x_1020; -x_316 = x_892; +x_293 = x_857; +x_294 = x_897; +x_295 = x_1020; +x_296 = x_896; +x_297 = x_968; +x_298 = x_864; +x_299 = x_928; +x_300 = x_865; +x_301 = x_866; +x_302 = x_914; +x_303 = x_869; +x_304 = x_912; +x_305 = x_1019; +x_306 = x_950; +x_307 = x_892; +x_308 = x_888; +x_309 = x_872; +x_310 = x_966; +x_311 = x_897; +x_312 = x_927; +x_313 = x_858; +x_314 = x_886; +x_315 = x_951; +x_316 = x_889; x_317 = x_861; -x_318 = x_863; -x_319 = x_864; -x_320 = x_950; -x_321 = x_866; -x_322 = x_870; -x_323 = x_871; -x_324 = x_897; -x_325 = x_928; +x_318 = x_862; +x_319 = x_863; +x_320 = x_910; +x_321 = x_882; +x_322 = x_867; +x_323 = x_868; +x_324 = x_870; +x_325 = x_871; x_326 = x_1022; goto block_335; } @@ -7682,39 +7682,39 @@ x_1025 = lean_ctor_get(x_1023, 0); x_1026 = 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__2_spec__2___lam__0(x_4, x_1025); lean_ctor_set(x_1023, 0, x_1026); x_1027 = l_Lean_MacroScopesView_review(x_1023); -x_293 = x_1019; -x_294 = x_857; -x_295 = x_951; -x_296 = x_886; -x_297 = x_858; -x_298 = x_910; -x_299 = x_927; -x_300 = x_882; -x_301 = x_862; -x_302 = x_912; -x_303 = x_865; -x_304 = x_867; -x_305 = x_888; -x_306 = x_869; -x_307 = x_914; -x_308 = x_897; -x_309 = x_889; -x_310 = x_968; -x_311 = x_896; -x_312 = x_859; -x_313 = x_966; -x_314 = x_860; -x_315 = x_1020; -x_316 = x_892; +x_293 = x_857; +x_294 = x_897; +x_295 = x_1020; +x_296 = x_896; +x_297 = x_968; +x_298 = x_864; +x_299 = x_928; +x_300 = x_865; +x_301 = x_866; +x_302 = x_914; +x_303 = x_869; +x_304 = x_912; +x_305 = x_1019; +x_306 = x_950; +x_307 = x_892; +x_308 = x_888; +x_309 = x_872; +x_310 = x_966; +x_311 = x_897; +x_312 = x_927; +x_313 = x_858; +x_314 = x_886; +x_315 = x_951; +x_316 = x_889; x_317 = x_861; -x_318 = x_863; -x_319 = x_864; -x_320 = x_950; -x_321 = x_866; -x_322 = x_870; -x_323 = x_871; -x_324 = x_897; -x_325 = x_928; +x_318 = x_862; +x_319 = x_863; +x_320 = x_910; +x_321 = x_882; +x_322 = x_867; +x_323 = x_868; +x_324 = x_870; +x_325 = x_871; x_326 = x_1027; goto block_335; } @@ -7737,39 +7737,39 @@ lean_ctor_set(x_1033, 1, x_1029); lean_ctor_set(x_1033, 2, x_1030); lean_ctor_set(x_1033, 3, x_1031); x_1034 = l_Lean_MacroScopesView_review(x_1033); -x_293 = x_1019; -x_294 = x_857; -x_295 = x_951; -x_296 = x_886; -x_297 = x_858; -x_298 = x_910; -x_299 = x_927; -x_300 = x_882; -x_301 = x_862; -x_302 = x_912; -x_303 = x_865; -x_304 = x_867; -x_305 = x_888; -x_306 = x_869; -x_307 = x_914; -x_308 = x_897; -x_309 = x_889; -x_310 = x_968; -x_311 = x_896; -x_312 = x_859; -x_313 = x_966; -x_314 = x_860; -x_315 = x_1020; -x_316 = x_892; +x_293 = x_857; +x_294 = x_897; +x_295 = x_1020; +x_296 = x_896; +x_297 = x_968; +x_298 = x_864; +x_299 = x_928; +x_300 = x_865; +x_301 = x_866; +x_302 = x_914; +x_303 = x_869; +x_304 = x_912; +x_305 = x_1019; +x_306 = x_950; +x_307 = x_892; +x_308 = x_888; +x_309 = x_872; +x_310 = x_966; +x_311 = x_897; +x_312 = x_927; +x_313 = x_858; +x_314 = x_886; +x_315 = x_951; +x_316 = x_889; x_317 = x_861; -x_318 = x_863; -x_319 = x_864; -x_320 = x_950; -x_321 = x_866; -x_322 = x_870; -x_323 = x_871; -x_324 = x_897; -x_325 = x_928; +x_318 = x_862; +x_319 = x_863; +x_320 = x_910; +x_321 = x_882; +x_322 = x_867; +x_323 = x_868; +x_324 = x_870; +x_325 = x_871; x_326 = x_1034; goto block_335; } @@ -7785,39 +7785,39 @@ if (x_1035 == 0) { lean_object* x_1036; x_1036 = 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__4_spec__4___lam__1(x_4, x_856); -x_813 = x_1019; -x_814 = x_857; -x_815 = x_951; -x_816 = x_858; -x_817 = x_886; -x_818 = x_910; -x_819 = x_927; -x_820 = x_882; -x_821 = x_862; -x_822 = x_912; -x_823 = x_865; -x_824 = x_867; -x_825 = x_888; -x_826 = x_869; -x_827 = x_914; -x_828 = x_897; -x_829 = x_889; -x_830 = x_968; -x_831 = x_896; -x_832 = x_859; -x_833 = x_966; -x_834 = x_860; -x_835 = x_1020; -x_836 = x_892; +x_813 = x_857; +x_814 = x_897; +x_815 = x_1020; +x_816 = x_896; +x_817 = x_968; +x_818 = x_864; +x_819 = x_928; +x_820 = x_865; +x_821 = x_866; +x_822 = x_914; +x_823 = x_869; +x_824 = x_912; +x_825 = x_1019; +x_826 = x_950; +x_827 = x_892; +x_828 = x_888; +x_829 = x_872; +x_830 = x_966; +x_831 = x_897; +x_832 = x_927; +x_833 = x_858; +x_834 = x_886; +x_835 = x_951; +x_836 = x_889; x_837 = x_861; -x_838 = x_863; -x_839 = x_864; -x_840 = x_950; -x_841 = x_866; -x_842 = x_870; -x_843 = x_871; -x_844 = x_928; -x_845 = x_897; +x_838 = x_862; +x_839 = x_863; +x_840 = x_910; +x_841 = x_882; +x_842 = x_868; +x_843 = x_867; +x_844 = x_870; +x_845 = x_871; x_846 = x_1036; goto block_855; } @@ -7833,39 +7833,39 @@ x_1039 = lean_ctor_get(x_1037, 0); x_1040 = 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__4_spec__4___lam__1(x_4, x_1039); lean_ctor_set(x_1037, 0, x_1040); x_1041 = l_Lean_MacroScopesView_review(x_1037); -x_813 = x_1019; -x_814 = x_857; -x_815 = x_951; -x_816 = x_858; -x_817 = x_886; -x_818 = x_910; -x_819 = x_927; -x_820 = x_882; -x_821 = x_862; -x_822 = x_912; -x_823 = x_865; -x_824 = x_867; -x_825 = x_888; -x_826 = x_869; -x_827 = x_914; -x_828 = x_897; -x_829 = x_889; -x_830 = x_968; -x_831 = x_896; -x_832 = x_859; -x_833 = x_966; -x_834 = x_860; -x_835 = x_1020; -x_836 = x_892; +x_813 = x_857; +x_814 = x_897; +x_815 = x_1020; +x_816 = x_896; +x_817 = x_968; +x_818 = x_864; +x_819 = x_928; +x_820 = x_865; +x_821 = x_866; +x_822 = x_914; +x_823 = x_869; +x_824 = x_912; +x_825 = x_1019; +x_826 = x_950; +x_827 = x_892; +x_828 = x_888; +x_829 = x_872; +x_830 = x_966; +x_831 = x_897; +x_832 = x_927; +x_833 = x_858; +x_834 = x_886; +x_835 = x_951; +x_836 = x_889; x_837 = x_861; -x_838 = x_863; -x_839 = x_864; -x_840 = x_950; -x_841 = x_866; -x_842 = x_870; -x_843 = x_871; -x_844 = x_928; -x_845 = x_897; +x_838 = x_862; +x_839 = x_863; +x_840 = x_910; +x_841 = x_882; +x_842 = x_868; +x_843 = x_867; +x_844 = x_870; +x_845 = x_871; x_846 = x_1041; goto block_855; } @@ -7888,39 +7888,39 @@ lean_ctor_set(x_1047, 1, x_1043); lean_ctor_set(x_1047, 2, x_1044); lean_ctor_set(x_1047, 3, x_1045); x_1048 = l_Lean_MacroScopesView_review(x_1047); -x_813 = x_1019; -x_814 = x_857; -x_815 = x_951; -x_816 = x_858; -x_817 = x_886; -x_818 = x_910; -x_819 = x_927; -x_820 = x_882; -x_821 = x_862; -x_822 = x_912; -x_823 = x_865; -x_824 = x_867; -x_825 = x_888; -x_826 = x_869; -x_827 = x_914; -x_828 = x_897; -x_829 = x_889; -x_830 = x_968; -x_831 = x_896; -x_832 = x_859; -x_833 = x_966; -x_834 = x_860; -x_835 = x_1020; -x_836 = x_892; +x_813 = x_857; +x_814 = x_897; +x_815 = x_1020; +x_816 = x_896; +x_817 = x_968; +x_818 = x_864; +x_819 = x_928; +x_820 = x_865; +x_821 = x_866; +x_822 = x_914; +x_823 = x_869; +x_824 = x_912; +x_825 = x_1019; +x_826 = x_950; +x_827 = x_892; +x_828 = x_888; +x_829 = x_872; +x_830 = x_966; +x_831 = x_897; +x_832 = x_927; +x_833 = x_858; +x_834 = x_886; +x_835 = x_951; +x_836 = x_889; x_837 = x_861; -x_838 = x_863; -x_839 = x_864; -x_840 = x_950; -x_841 = x_866; -x_842 = x_870; -x_843 = x_871; -x_844 = x_928; -x_845 = x_897; +x_838 = x_862; +x_839 = x_863; +x_840 = x_910; +x_841 = x_882; +x_842 = x_868; +x_843 = x_867; +x_844 = x_870; +x_845 = x_871; x_846 = x_1048; goto block_855; } @@ -7955,27 +7955,27 @@ if (lean_obj_tag(x_3) == 0) lean_object* x_1067; x_1067 = 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__2_spec__2___closed__53; lean_inc(x_1054); -lean_inc(x_1055); -lean_inc(x_1056); lean_inc(x_1051); -lean_inc(x_1052); +lean_inc(x_1055); lean_inc(x_1053); +lean_inc(x_1052); +lean_inc(x_1056); x_857 = x_1064; -x_858 = x_1053; -x_859 = x_1052; -x_860 = x_1063; -x_861 = x_1061; -x_862 = x_1057; -x_863 = x_1060; -x_864 = x_1051; +x_858 = x_1062; +x_859 = x_1066; +x_860 = x_1058; +x_861 = x_1065; +x_862 = x_1063; +x_863 = x_1059; +x_864 = x_1060; x_865 = x_1056; -x_866 = x_1055; -x_867 = x_1062; -x_868 = x_1066; -x_869 = x_1059; -x_870 = x_1054; -x_871 = x_1065; -x_872 = x_1058; +x_866 = x_1052; +x_867 = x_1061; +x_868 = x_1053; +x_869 = x_1055; +x_870 = x_1051; +x_871 = x_1057; +x_872 = x_1054; x_873 = x_1067; goto block_1049; } @@ -7986,27 +7986,27 @@ x_1068 = lean_ctor_get(x_3, 0); lean_inc(x_1068); x_1069 = l_Array_mkArray1___redArg(x_1068); lean_inc(x_1054); -lean_inc(x_1055); -lean_inc(x_1056); lean_inc(x_1051); -lean_inc(x_1052); +lean_inc(x_1055); lean_inc(x_1053); +lean_inc(x_1052); +lean_inc(x_1056); x_857 = x_1064; -x_858 = x_1053; -x_859 = x_1052; -x_860 = x_1063; -x_861 = x_1061; -x_862 = x_1057; -x_863 = x_1060; -x_864 = x_1051; +x_858 = x_1062; +x_859 = x_1066; +x_860 = x_1058; +x_861 = x_1065; +x_862 = x_1063; +x_863 = x_1059; +x_864 = x_1060; x_865 = x_1056; -x_866 = x_1055; -x_867 = x_1062; -x_868 = x_1066; -x_869 = x_1059; -x_870 = x_1054; -x_871 = x_1065; -x_872 = x_1058; +x_866 = x_1052; +x_867 = x_1061; +x_868 = x_1053; +x_869 = x_1055; +x_870 = x_1051; +x_871 = x_1057; +x_872 = x_1054; x_873 = x_1069; goto block_1049; } @@ -8119,19 +8119,19 @@ goto block_1070; block_292: { lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = l_Lean_replaceRef(x_19, x_55); -lean_dec(x_55); +x_65 = l_Lean_replaceRef(x_19, x_61); +lean_dec(x_61); lean_inc(x_65); +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); x_66 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_66, 0, x_48); -lean_ctor_set(x_66, 1, x_47); -lean_ctor_set(x_66, 2, x_46); -lean_ctor_set(x_66, 3, x_42); -lean_ctor_set(x_66, 4, x_31); +lean_ctor_set(x_66, 0, x_28); +lean_ctor_set(x_66, 1, x_46); +lean_ctor_set(x_66, 2, x_51); +lean_ctor_set(x_66, 3, x_27); +lean_ctor_set(x_66, 4, x_32); lean_ctor_set(x_66, 5, x_65); -x_67 = 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__2_spec__2___lam__1(x_17, x_65, x_66, x_52); +x_67 = 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__2_spec__2___lam__1(x_17, x_65, x_66, x_30); lean_dec_ref(x_66); lean_dec(x_65); x_68 = !lean_is_exclusive(x_67); @@ -8140,139 +8140,139 @@ if (x_68 == 0) lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; x_69 = lean_ctor_get(x_67, 0); x_70 = lean_ctor_get(x_67, 1); -lean_inc_ref(x_35); -x_71 = l_Array_append___redArg(x_35, x_64); +lean_inc_ref(x_48); +x_71 = l_Array_append___redArg(x_48, x_64); lean_dec_ref(x_64); -lean_inc(x_40); -lean_inc(x_60); +lean_inc(x_54); +lean_inc(x_34); x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_60); -lean_ctor_set(x_72, 1, x_40); +lean_ctor_set(x_72, 0, x_34); +lean_ctor_set(x_72, 1, x_54); lean_ctor_set(x_72, 2, x_71); -lean_inc_n(x_51, 6); -lean_inc(x_60); -x_73 = l_Lean_Syntax_node7(x_60, x_36, x_51, x_51, x_72, x_51, x_51, x_51, x_51); +lean_inc_n(x_50, 6); +lean_inc(x_34); +x_73 = l_Lean_Syntax_node7(x_34, x_29, x_50, x_50, x_72, x_50, x_50, x_50, x_50); x_74 = 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__2_spec__2___closed__5; -lean_inc_ref(x_32); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_75 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_74); +lean_inc_ref(x_55); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_75 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_74); x_76 = 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__2_spec__2___closed__6; -lean_inc_ref(x_38); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_77 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_76); -lean_inc(x_51); -lean_inc(x_60); -x_78 = l_Lean_Syntax_node1(x_60, x_77, x_51); -lean_inc(x_60); +lean_inc_ref(x_47); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_77 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_76); +lean_inc(x_50); +lean_inc(x_34); +x_78 = l_Lean_Syntax_node1(x_34, x_77, x_50); +lean_inc(x_34); lean_ctor_set_tag(x_67, 2); lean_ctor_set(x_67, 1, x_74); -lean_ctor_set(x_67, 0, x_60); -lean_inc(x_51); -lean_inc(x_60); -x_79 = l_Lean_Syntax_node2(x_60, x_34, x_41, x_51); -lean_inc(x_40); -lean_inc(x_60); -x_80 = l_Lean_Syntax_node1(x_60, x_40, x_79); +lean_ctor_set(x_67, 0, x_34); +lean_inc(x_50); +lean_inc(x_34); +x_79 = l_Lean_Syntax_node2(x_34, x_44, x_57, x_50); +lean_inc(x_54); +lean_inc(x_34); +x_80 = l_Lean_Syntax_node1(x_34, x_54, x_79); x_81 = l_Lake_configField___closed__19; -lean_inc_ref(x_32); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_82 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_81); +lean_inc_ref(x_55); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_82 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_81); lean_inc_ref(x_45); -lean_inc(x_60); +lean_inc(x_34); x_83 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_83, 0, x_60); +lean_ctor_set(x_83, 0, x_34); lean_ctor_set(x_83, 1, x_45); x_84 = 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__2_spec__2___closed__10; x_85 = 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__2_spec__2___closed__11; x_86 = 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__2_spec__2___closed__12; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_87 = l_Lean_addMacroScope(x_47, x_86, x_46); -lean_inc_ref(x_63); -x_88 = l_Lean_Name_mkStr2(x_63, x_84); -lean_inc(x_54); +x_87 = l_Lean_addMacroScope(x_46, x_86, x_51); +lean_inc_ref(x_53); +x_88 = l_Lean_Name_mkStr2(x_53, x_84); +lean_inc(x_38); lean_inc(x_88); x_89 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_54); +lean_ctor_set(x_89, 1, x_38); x_90 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_90, 0, x_88); -lean_inc(x_62); +lean_inc(x_49); x_91 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_62); +lean_ctor_set(x_91, 1, x_49); x_92 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_92, 0, x_89); lean_ctor_set(x_92, 1, x_91); -lean_inc(x_60); +lean_inc(x_34); x_93 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_93, 0, x_60); +lean_ctor_set(x_93, 0, x_34); lean_ctor_set(x_93, 1, x_85); lean_ctor_set(x_93, 2, x_87); lean_ctor_set(x_93, 3, x_92); lean_inc(x_24); -lean_inc(x_49); -lean_inc(x_1); lean_inc(x_40); -lean_inc(x_60); -x_94 = l_Lean_Syntax_node3(x_60, x_40, x_1, x_49, x_24); -lean_inc(x_60); -x_95 = l_Lean_Syntax_node2(x_60, x_43, x_93, x_94); -lean_inc(x_60); -x_96 = l_Lean_Syntax_node2(x_60, x_59, x_83, x_95); -lean_inc(x_51); -lean_inc(x_60); -x_97 = l_Lean_Syntax_node2(x_60, x_82, x_51, x_96); -x_98 = 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__2_spec__2___closed__18; -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_99 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_98); -lean_inc_ref(x_56); -lean_inc(x_60); +lean_inc(x_1); +lean_inc(x_54); +lean_inc(x_34); +x_94 = l_Lean_Syntax_node3(x_34, x_54, x_1, x_40, x_24); +lean_inc(x_34); +x_95 = l_Lean_Syntax_node2(x_34, x_41, x_93, x_94); +lean_inc(x_34); +x_96 = l_Lean_Syntax_node2(x_34, x_62, x_83, x_95); +lean_inc(x_50); +lean_inc(x_34); +x_97 = l_Lean_Syntax_node2(x_34, x_82, x_50, x_96); +x_98 = 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__2_spec__2___closed__18; +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_99 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_98); +lean_inc_ref(x_60); +lean_inc(x_34); x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_60); -lean_ctor_set(x_100, 1, x_56); +lean_ctor_set(x_100, 0, x_34); +lean_ctor_set(x_100, 1, x_60); x_101 = 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__2_spec__2___closed__20; -lean_inc_ref(x_38); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_102 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_101); +lean_inc_ref(x_47); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_102 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_101); x_103 = 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__2_spec__2___closed__21; -lean_inc(x_60); +lean_inc(x_34); x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_60); +lean_ctor_set(x_104, 0, x_34); lean_ctor_set(x_104, 1, x_103); -lean_inc(x_44); -lean_inc(x_40); -lean_inc(x_60); -x_105 = l_Lean_Syntax_node1(x_60, x_40, x_44); +lean_inc(x_56); +lean_inc(x_54); +lean_inc(x_34); +x_105 = l_Lean_Syntax_node1(x_34, x_54, x_56); x_106 = 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__2_spec__2___closed__22; -lean_inc(x_60); +lean_inc(x_34); x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_60); +lean_ctor_set(x_107, 0, x_34); lean_ctor_set(x_107, 1, x_106); -lean_inc(x_60); -x_108 = l_Lean_Syntax_node3(x_60, x_102, x_104, x_105, x_107); +lean_inc(x_34); +x_108 = l_Lean_Syntax_node3(x_34, x_102, x_104, x_105, x_107); x_109 = 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__2_spec__2___closed__23; x_110 = 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__2_spec__2___closed__24; -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_111 = l_Lean_Name_mkStr4(x_53, x_33, x_109, x_110); -lean_inc_n(x_51, 2); -lean_inc(x_60); -x_112 = l_Lean_Syntax_node2(x_60, x_111, x_51, x_51); -lean_inc(x_51); -lean_inc(x_60); -x_113 = l_Lean_Syntax_node4(x_60, x_99, x_100, x_108, x_112, x_51); -lean_inc(x_60); -x_114 = l_Lean_Syntax_node6(x_60, x_75, x_78, x_67, x_51, x_80, x_97, x_113); -x_115 = l_Lean_Syntax_node2(x_60, x_39, x_73, x_114); -x_116 = lean_array_push(x_30, x_115); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_111 = l_Lean_Name_mkStr4(x_58, x_35, x_109, x_110); +lean_inc_n(x_50, 2); +lean_inc(x_34); +x_112 = l_Lean_Syntax_node2(x_34, x_111, x_50, x_50); +lean_inc(x_50); +lean_inc(x_34); +x_113 = l_Lean_Syntax_node4(x_34, x_99, x_100, x_108, x_112, x_50); +lean_inc(x_34); +x_114 = l_Lean_Syntax_node6(x_34, x_75, x_78, x_67, x_50, x_80, x_97, x_113); +x_115 = l_Lean_Syntax_node2(x_34, x_59, x_73, x_114); +x_116 = lean_array_push(x_37, x_115); x_117 = 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__2_spec__2___closed__25; -x_118 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_117); +x_118 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_117); x_119 = 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__2_spec__2___closed__26; lean_inc(x_69); x_120 = lean_alloc_ctor(2, 2, 0); @@ -8280,113 +8280,113 @@ lean_ctor_set(x_120, 0, x_69); lean_ctor_set(x_120, 1, x_119); x_121 = 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__2_spec__2___closed__28; x_122 = 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__2_spec__2___closed__29; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_123 = l_Lean_addMacroScope(x_47, x_122, x_46); -lean_inc(x_62); +x_123 = l_Lean_addMacroScope(x_46, x_122, x_51); +lean_inc(x_49); lean_inc(x_69); x_124 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_124, 0, x_69); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_62); +lean_ctor_set(x_124, 3, x_49); lean_inc(x_69); x_125 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_125, 0, x_69); -lean_ctor_set(x_125, 1, x_58); -lean_inc(x_40); +lean_ctor_set(x_125, 1, x_36); +lean_inc(x_54); lean_inc(x_69); x_126 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_126, 0, x_69); -lean_ctor_set(x_126, 1, x_40); -lean_ctor_set(x_126, 2, x_35); +lean_ctor_set(x_126, 1, x_54); +lean_ctor_set(x_126, 2, x_48); x_127 = 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__2_spec__2___closed__34; x_128 = 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__2_spec__2___closed__35; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_129 = l_Lean_addMacroScope(x_47, x_128, x_46); -lean_inc(x_62); +x_129 = l_Lean_addMacroScope(x_46, x_128, x_51); +lean_inc(x_49); lean_inc(x_69); x_130 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_130, 0, x_69); lean_ctor_set(x_130, 1, x_127); lean_ctor_set(x_130, 2, x_129); -lean_ctor_set(x_130, 3, x_62); +lean_ctor_set(x_130, 3, x_49); lean_inc_ref(x_126); -lean_inc(x_37); +lean_inc(x_43); lean_inc(x_69); -x_131 = l_Lean_Syntax_node2(x_69, x_37, x_130, x_126); +x_131 = l_Lean_Syntax_node2(x_69, x_43, x_130, x_126); lean_inc(x_69); x_132 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_132, 0, x_69); -lean_ctor_set(x_132, 1, x_56); -lean_inc(x_49); +lean_ctor_set(x_132, 1, x_60); +lean_inc(x_40); lean_inc_ref(x_126); lean_inc_ref(x_132); -lean_inc(x_27); +lean_inc(x_63); lean_inc(x_69); -x_133 = l_Lean_Syntax_node3(x_69, x_27, x_132, x_126, x_49); +x_133 = l_Lean_Syntax_node3(x_69, x_63, x_132, x_126, x_40); lean_inc_ref_n(x_126, 2); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_69); -x_134 = l_Lean_Syntax_node3(x_69, x_40, x_126, x_126, x_133); +x_134 = l_Lean_Syntax_node3(x_69, x_54, x_126, x_126, x_133); lean_inc(x_134); -lean_inc(x_29); +lean_inc(x_33); lean_inc(x_69); -x_135 = l_Lean_Syntax_node2(x_69, x_29, x_131, x_134); +x_135 = l_Lean_Syntax_node2(x_69, x_33, x_131, x_134); x_136 = 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__2_spec__2___closed__38; x_137 = 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__2_spec__2___closed__39; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_138 = l_Lean_addMacroScope(x_47, x_137, x_46); -lean_inc(x_62); +x_138 = l_Lean_addMacroScope(x_46, x_137, x_51); +lean_inc(x_49); lean_inc(x_69); x_139 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_139, 0, x_69); lean_ctor_set(x_139, 1, x_136); lean_ctor_set(x_139, 2, x_138); -lean_ctor_set(x_139, 3, x_62); +lean_ctor_set(x_139, 3, x_49); lean_inc_ref(x_126); -lean_inc(x_37); +lean_inc(x_43); lean_inc(x_69); -x_140 = l_Lean_Syntax_node2(x_69, x_37, x_139, x_126); -lean_inc(x_29); +x_140 = l_Lean_Syntax_node2(x_69, x_43, x_139, x_126); +lean_inc(x_33); lean_inc(x_69); -x_141 = l_Lean_Syntax_node2(x_69, x_29, x_140, x_134); +x_141 = l_Lean_Syntax_node2(x_69, x_33, x_140, x_134); x_142 = 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__2_spec__2___closed__41; x_143 = 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__2_spec__2___closed__42; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_144 = l_Lean_addMacroScope(x_47, x_143, x_46); -lean_inc(x_62); +x_144 = l_Lean_addMacroScope(x_46, x_143, x_51); +lean_inc(x_49); lean_inc(x_69); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_69); lean_ctor_set(x_145, 1, x_142); lean_ctor_set(x_145, 2, x_144); -lean_ctor_set(x_145, 3, x_62); +lean_ctor_set(x_145, 3, x_49); lean_inc_ref(x_126); lean_inc(x_69); -x_146 = l_Lean_Syntax_node2(x_69, x_37, x_145, x_126); +x_146 = l_Lean_Syntax_node2(x_69, x_43, x_145, x_126); x_147 = 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__2_spec__2___closed__49; lean_inc_ref(x_126); lean_inc(x_69); -x_148 = l_Lean_Syntax_node3(x_69, x_27, x_132, x_126, x_147); +x_148 = l_Lean_Syntax_node3(x_69, x_63, x_132, x_126, x_147); lean_inc_ref_n(x_126, 2); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_69); -x_149 = l_Lean_Syntax_node3(x_69, x_40, x_126, x_126, x_148); +x_149 = l_Lean_Syntax_node3(x_69, x_54, x_126, x_126, x_148); lean_inc(x_69); -x_150 = l_Lean_Syntax_node2(x_69, x_29, x_146, x_149); +x_150 = l_Lean_Syntax_node2(x_69, x_33, x_146, x_149); lean_inc_ref_n(x_126, 3); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_69); -x_151 = l_Lean_Syntax_node6(x_69, x_40, x_135, x_126, x_141, x_126, x_150, x_126); +x_151 = l_Lean_Syntax_node6(x_69, x_54, x_135, x_126, x_141, x_126, x_150, x_126); lean_inc(x_69); -x_152 = l_Lean_Syntax_node1(x_69, x_28, x_151); +x_152 = l_Lean_Syntax_node1(x_69, x_42, x_151); lean_inc_ref(x_126); lean_inc(x_69); -x_153 = l_Lean_Syntax_node1(x_69, x_50, x_126); +x_153 = l_Lean_Syntax_node1(x_69, x_39, x_126); lean_inc(x_69); x_154 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_154, 0, x_69); @@ -8394,17 +8394,17 @@ lean_ctor_set(x_154, 1, x_45); x_155 = 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__2_spec__2___closed__1; x_156 = 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__2_spec__2___closed__2; x_157 = 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__2_spec__2___closed__3; -x_158 = l_Lean_addMacroScope(x_47, x_157, x_46); -x_159 = l_Lean_Name_mkStr2(x_63, x_155); +x_158 = l_Lean_addMacroScope(x_46, x_157, x_51); +x_159 = l_Lean_Name_mkStr2(x_53, x_155); lean_inc(x_159); x_160 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_54); +lean_ctor_set(x_160, 1, x_38); x_161 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_161, 0, x_159); x_162 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_62); +lean_ctor_set(x_162, 1, x_49); x_163 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_163, 0, x_160); lean_ctor_set(x_163, 1, x_162); @@ -8414,17 +8414,17 @@ lean_ctor_set(x_164, 0, x_69); lean_ctor_set(x_164, 1, x_156); lean_ctor_set(x_164, 2, x_158); lean_ctor_set(x_164, 3, x_163); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_69); -x_165 = l_Lean_Syntax_node2(x_69, x_40, x_154, x_164); +x_165 = l_Lean_Syntax_node2(x_69, x_54, x_154, x_164); lean_inc(x_69); x_166 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_166, 0, x_69); -lean_ctor_set(x_166, 1, x_61); +lean_ctor_set(x_166, 1, x_52); lean_inc(x_69); -x_167 = l_Lean_Syntax_node6(x_69, x_57, x_125, x_126, x_152, x_153, x_165, x_166); +x_167 = l_Lean_Syntax_node6(x_69, x_31, x_125, x_126, x_152, x_153, x_165, x_166); lean_inc(x_69); -x_168 = l_Lean_Syntax_node1(x_69, x_40, x_167); +x_168 = l_Lean_Syntax_node1(x_69, x_54, x_167); x_169 = l_Lean_Syntax_node4(x_69, x_118, x_19, x_120, x_124, x_168); if (lean_is_scalar(x_20)) { x_170 = lean_alloc_ctor(0, 2, 0); @@ -8439,8 +8439,8 @@ x_173 = lean_nat_dec_lt(x_171, x_172); if (x_173 == 0) { lean_dec(x_172); -lean_dec(x_49); -lean_dec(x_44); +lean_dec(x_56); +lean_dec(x_40); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_170; @@ -8454,8 +8454,8 @@ x_174 = lean_nat_dec_le(x_172, x_172); if (x_174 == 0) { lean_dec(x_172); -lean_dec(x_49); -lean_dec(x_44); +lean_dec(x_56); +lean_dec(x_40); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_170; @@ -8471,7 +8471,7 @@ lean_dec(x_172); lean_inc_ref(x_9); lean_inc(x_3); lean_inc(x_1); -x_177 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_44, x_49, x_3, x_4, x_23, x_175, x_176, x_170, x_9, x_70); +x_177 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_56, x_40, x_3, x_4, x_23, x_175, x_176, x_170, x_9, x_70); lean_dec_ref(x_23); x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); @@ -8492,139 +8492,139 @@ x_181 = lean_ctor_get(x_67, 1); lean_inc(x_181); lean_inc(x_180); lean_dec(x_67); -lean_inc_ref(x_35); -x_182 = l_Array_append___redArg(x_35, x_64); +lean_inc_ref(x_48); +x_182 = l_Array_append___redArg(x_48, x_64); lean_dec_ref(x_64); -lean_inc(x_40); -lean_inc(x_60); +lean_inc(x_54); +lean_inc(x_34); x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_60); -lean_ctor_set(x_183, 1, x_40); +lean_ctor_set(x_183, 0, x_34); +lean_ctor_set(x_183, 1, x_54); lean_ctor_set(x_183, 2, x_182); -lean_inc_n(x_51, 6); -lean_inc(x_60); -x_184 = l_Lean_Syntax_node7(x_60, x_36, x_51, x_51, x_183, x_51, x_51, x_51, x_51); +lean_inc_n(x_50, 6); +lean_inc(x_34); +x_184 = l_Lean_Syntax_node7(x_34, x_29, x_50, x_50, x_183, x_50, x_50, x_50, x_50); x_185 = 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__2_spec__2___closed__5; -lean_inc_ref(x_32); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_186 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_185); +lean_inc_ref(x_55); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_186 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_185); x_187 = 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__2_spec__2___closed__6; -lean_inc_ref(x_38); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_188 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_187); -lean_inc(x_51); -lean_inc(x_60); -x_189 = l_Lean_Syntax_node1(x_60, x_188, x_51); -lean_inc(x_60); +lean_inc_ref(x_47); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_188 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_187); +lean_inc(x_50); +lean_inc(x_34); +x_189 = l_Lean_Syntax_node1(x_34, x_188, x_50); +lean_inc(x_34); x_190 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_190, 0, x_60); +lean_ctor_set(x_190, 0, x_34); lean_ctor_set(x_190, 1, x_185); -lean_inc(x_51); -lean_inc(x_60); -x_191 = l_Lean_Syntax_node2(x_60, x_34, x_41, x_51); -lean_inc(x_40); -lean_inc(x_60); -x_192 = l_Lean_Syntax_node1(x_60, x_40, x_191); +lean_inc(x_50); +lean_inc(x_34); +x_191 = l_Lean_Syntax_node2(x_34, x_44, x_57, x_50); +lean_inc(x_54); +lean_inc(x_34); +x_192 = l_Lean_Syntax_node1(x_34, x_54, x_191); x_193 = l_Lake_configField___closed__19; -lean_inc_ref(x_32); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_194 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_193); +lean_inc_ref(x_55); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_194 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_193); lean_inc_ref(x_45); -lean_inc(x_60); +lean_inc(x_34); x_195 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_195, 0, x_60); +lean_ctor_set(x_195, 0, x_34); lean_ctor_set(x_195, 1, x_45); x_196 = 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__2_spec__2___closed__10; x_197 = 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__2_spec__2___closed__11; x_198 = 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__2_spec__2___closed__12; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_199 = l_Lean_addMacroScope(x_47, x_198, x_46); -lean_inc_ref(x_63); -x_200 = l_Lean_Name_mkStr2(x_63, x_196); -lean_inc(x_54); +x_199 = l_Lean_addMacroScope(x_46, x_198, x_51); +lean_inc_ref(x_53); +x_200 = l_Lean_Name_mkStr2(x_53, x_196); +lean_inc(x_38); lean_inc(x_200); x_201 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_54); +lean_ctor_set(x_201, 1, x_38); x_202 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_202, 0, x_200); -lean_inc(x_62); +lean_inc(x_49); x_203 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_62); +lean_ctor_set(x_203, 1, x_49); x_204 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_204, 0, x_201); lean_ctor_set(x_204, 1, x_203); -lean_inc(x_60); +lean_inc(x_34); x_205 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_205, 0, x_60); +lean_ctor_set(x_205, 0, x_34); lean_ctor_set(x_205, 1, x_197); lean_ctor_set(x_205, 2, x_199); lean_ctor_set(x_205, 3, x_204); lean_inc(x_24); -lean_inc(x_49); -lean_inc(x_1); lean_inc(x_40); -lean_inc(x_60); -x_206 = l_Lean_Syntax_node3(x_60, x_40, x_1, x_49, x_24); -lean_inc(x_60); -x_207 = l_Lean_Syntax_node2(x_60, x_43, x_205, x_206); -lean_inc(x_60); -x_208 = l_Lean_Syntax_node2(x_60, x_59, x_195, x_207); -lean_inc(x_51); -lean_inc(x_60); -x_209 = l_Lean_Syntax_node2(x_60, x_194, x_51, x_208); +lean_inc(x_1); +lean_inc(x_54); +lean_inc(x_34); +x_206 = l_Lean_Syntax_node3(x_34, x_54, x_1, x_40, x_24); +lean_inc(x_34); +x_207 = l_Lean_Syntax_node2(x_34, x_41, x_205, x_206); +lean_inc(x_34); +x_208 = l_Lean_Syntax_node2(x_34, x_62, x_195, x_207); +lean_inc(x_50); +lean_inc(x_34); +x_209 = l_Lean_Syntax_node2(x_34, x_194, x_50, x_208); x_210 = 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__2_spec__2___closed__18; -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_211 = l_Lean_Name_mkStr4(x_53, x_33, x_32, x_210); -lean_inc_ref(x_56); -lean_inc(x_60); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_211 = l_Lean_Name_mkStr4(x_58, x_35, x_55, x_210); +lean_inc_ref(x_60); +lean_inc(x_34); x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_60); -lean_ctor_set(x_212, 1, x_56); +lean_ctor_set(x_212, 0, x_34); +lean_ctor_set(x_212, 1, x_60); x_213 = 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__2_spec__2___closed__20; -lean_inc_ref(x_38); -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_214 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_213); +lean_inc_ref(x_47); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_214 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_213); x_215 = 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__2_spec__2___closed__21; -lean_inc(x_60); +lean_inc(x_34); x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_60); +lean_ctor_set(x_216, 0, x_34); lean_ctor_set(x_216, 1, x_215); -lean_inc(x_44); -lean_inc(x_40); -lean_inc(x_60); -x_217 = l_Lean_Syntax_node1(x_60, x_40, x_44); +lean_inc(x_56); +lean_inc(x_54); +lean_inc(x_34); +x_217 = l_Lean_Syntax_node1(x_34, x_54, x_56); x_218 = 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__2_spec__2___closed__22; -lean_inc(x_60); +lean_inc(x_34); x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_60); +lean_ctor_set(x_219, 0, x_34); lean_ctor_set(x_219, 1, x_218); -lean_inc(x_60); -x_220 = l_Lean_Syntax_node3(x_60, x_214, x_216, x_217, x_219); +lean_inc(x_34); +x_220 = l_Lean_Syntax_node3(x_34, x_214, x_216, x_217, x_219); x_221 = 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__2_spec__2___closed__23; x_222 = 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__2_spec__2___closed__24; -lean_inc_ref(x_33); -lean_inc_ref(x_53); -x_223 = l_Lean_Name_mkStr4(x_53, x_33, x_221, x_222); -lean_inc_n(x_51, 2); -lean_inc(x_60); -x_224 = l_Lean_Syntax_node2(x_60, x_223, x_51, x_51); -lean_inc(x_51); -lean_inc(x_60); -x_225 = l_Lean_Syntax_node4(x_60, x_211, x_212, x_220, x_224, x_51); -lean_inc(x_60); -x_226 = l_Lean_Syntax_node6(x_60, x_186, x_189, x_190, x_51, x_192, x_209, x_225); -x_227 = l_Lean_Syntax_node2(x_60, x_39, x_184, x_226); -x_228 = lean_array_push(x_30, x_227); +lean_inc_ref(x_35); +lean_inc_ref(x_58); +x_223 = l_Lean_Name_mkStr4(x_58, x_35, x_221, x_222); +lean_inc_n(x_50, 2); +lean_inc(x_34); +x_224 = l_Lean_Syntax_node2(x_34, x_223, x_50, x_50); +lean_inc(x_50); +lean_inc(x_34); +x_225 = l_Lean_Syntax_node4(x_34, x_211, x_212, x_220, x_224, x_50); +lean_inc(x_34); +x_226 = l_Lean_Syntax_node6(x_34, x_186, x_189, x_190, x_50, x_192, x_209, x_225); +x_227 = l_Lean_Syntax_node2(x_34, x_59, x_184, x_226); +x_228 = lean_array_push(x_37, x_227); x_229 = 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__2_spec__2___closed__25; -x_230 = l_Lean_Name_mkStr4(x_53, x_33, x_38, x_229); +x_230 = l_Lean_Name_mkStr4(x_58, x_35, x_47, x_229); x_231 = 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__2_spec__2___closed__26; lean_inc(x_180); x_232 = lean_alloc_ctor(2, 2, 0); @@ -8632,113 +8632,113 @@ lean_ctor_set(x_232, 0, x_180); lean_ctor_set(x_232, 1, x_231); x_233 = 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__2_spec__2___closed__28; x_234 = 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__2_spec__2___closed__29; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_235 = l_Lean_addMacroScope(x_47, x_234, x_46); -lean_inc(x_62); +x_235 = l_Lean_addMacroScope(x_46, x_234, x_51); +lean_inc(x_49); lean_inc(x_180); x_236 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_236, 0, x_180); lean_ctor_set(x_236, 1, x_233); lean_ctor_set(x_236, 2, x_235); -lean_ctor_set(x_236, 3, x_62); +lean_ctor_set(x_236, 3, x_49); lean_inc(x_180); x_237 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_237, 0, x_180); -lean_ctor_set(x_237, 1, x_58); -lean_inc(x_40); +lean_ctor_set(x_237, 1, x_36); +lean_inc(x_54); lean_inc(x_180); x_238 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_238, 0, x_180); -lean_ctor_set(x_238, 1, x_40); -lean_ctor_set(x_238, 2, x_35); +lean_ctor_set(x_238, 1, x_54); +lean_ctor_set(x_238, 2, x_48); x_239 = 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__2_spec__2___closed__34; x_240 = 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__2_spec__2___closed__35; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_241 = l_Lean_addMacroScope(x_47, x_240, x_46); -lean_inc(x_62); +x_241 = l_Lean_addMacroScope(x_46, x_240, x_51); +lean_inc(x_49); lean_inc(x_180); x_242 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_242, 0, x_180); lean_ctor_set(x_242, 1, x_239); lean_ctor_set(x_242, 2, x_241); -lean_ctor_set(x_242, 3, x_62); +lean_ctor_set(x_242, 3, x_49); lean_inc_ref(x_238); -lean_inc(x_37); +lean_inc(x_43); lean_inc(x_180); -x_243 = l_Lean_Syntax_node2(x_180, x_37, x_242, x_238); +x_243 = l_Lean_Syntax_node2(x_180, x_43, x_242, x_238); lean_inc(x_180); x_244 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_244, 0, x_180); -lean_ctor_set(x_244, 1, x_56); -lean_inc(x_49); +lean_ctor_set(x_244, 1, x_60); +lean_inc(x_40); lean_inc_ref(x_238); lean_inc_ref(x_244); -lean_inc(x_27); +lean_inc(x_63); lean_inc(x_180); -x_245 = l_Lean_Syntax_node3(x_180, x_27, x_244, x_238, x_49); +x_245 = l_Lean_Syntax_node3(x_180, x_63, x_244, x_238, x_40); lean_inc_ref_n(x_238, 2); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_180); -x_246 = l_Lean_Syntax_node3(x_180, x_40, x_238, x_238, x_245); +x_246 = l_Lean_Syntax_node3(x_180, x_54, x_238, x_238, x_245); lean_inc(x_246); -lean_inc(x_29); +lean_inc(x_33); lean_inc(x_180); -x_247 = l_Lean_Syntax_node2(x_180, x_29, x_243, x_246); +x_247 = l_Lean_Syntax_node2(x_180, x_33, x_243, x_246); x_248 = 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__2_spec__2___closed__38; x_249 = 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__2_spec__2___closed__39; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_250 = l_Lean_addMacroScope(x_47, x_249, x_46); -lean_inc(x_62); +x_250 = l_Lean_addMacroScope(x_46, x_249, x_51); +lean_inc(x_49); lean_inc(x_180); x_251 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_251, 0, x_180); lean_ctor_set(x_251, 1, x_248); lean_ctor_set(x_251, 2, x_250); -lean_ctor_set(x_251, 3, x_62); +lean_ctor_set(x_251, 3, x_49); lean_inc_ref(x_238); -lean_inc(x_37); +lean_inc(x_43); lean_inc(x_180); -x_252 = l_Lean_Syntax_node2(x_180, x_37, x_251, x_238); -lean_inc(x_29); +x_252 = l_Lean_Syntax_node2(x_180, x_43, x_251, x_238); +lean_inc(x_33); lean_inc(x_180); -x_253 = l_Lean_Syntax_node2(x_180, x_29, x_252, x_246); +x_253 = l_Lean_Syntax_node2(x_180, x_33, x_252, x_246); x_254 = 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__2_spec__2___closed__41; x_255 = 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__2_spec__2___closed__42; +lean_inc(x_51); lean_inc(x_46); -lean_inc(x_47); -x_256 = l_Lean_addMacroScope(x_47, x_255, x_46); -lean_inc(x_62); +x_256 = l_Lean_addMacroScope(x_46, x_255, x_51); +lean_inc(x_49); lean_inc(x_180); x_257 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_257, 0, x_180); lean_ctor_set(x_257, 1, x_254); lean_ctor_set(x_257, 2, x_256); -lean_ctor_set(x_257, 3, x_62); +lean_ctor_set(x_257, 3, x_49); lean_inc_ref(x_238); lean_inc(x_180); -x_258 = l_Lean_Syntax_node2(x_180, x_37, x_257, x_238); +x_258 = l_Lean_Syntax_node2(x_180, x_43, x_257, x_238); x_259 = 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__2_spec__2___closed__49; lean_inc_ref(x_238); lean_inc(x_180); -x_260 = l_Lean_Syntax_node3(x_180, x_27, x_244, x_238, x_259); +x_260 = l_Lean_Syntax_node3(x_180, x_63, x_244, x_238, x_259); lean_inc_ref_n(x_238, 2); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_180); -x_261 = l_Lean_Syntax_node3(x_180, x_40, x_238, x_238, x_260); +x_261 = l_Lean_Syntax_node3(x_180, x_54, x_238, x_238, x_260); lean_inc(x_180); -x_262 = l_Lean_Syntax_node2(x_180, x_29, x_258, x_261); +x_262 = l_Lean_Syntax_node2(x_180, x_33, x_258, x_261); lean_inc_ref_n(x_238, 3); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_180); -x_263 = l_Lean_Syntax_node6(x_180, x_40, x_247, x_238, x_253, x_238, x_262, x_238); +x_263 = l_Lean_Syntax_node6(x_180, x_54, x_247, x_238, x_253, x_238, x_262, x_238); lean_inc(x_180); -x_264 = l_Lean_Syntax_node1(x_180, x_28, x_263); +x_264 = l_Lean_Syntax_node1(x_180, x_42, x_263); lean_inc_ref(x_238); lean_inc(x_180); -x_265 = l_Lean_Syntax_node1(x_180, x_50, x_238); +x_265 = l_Lean_Syntax_node1(x_180, x_39, x_238); lean_inc(x_180); x_266 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_266, 0, x_180); @@ -8746,17 +8746,17 @@ lean_ctor_set(x_266, 1, x_45); x_267 = 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__2_spec__2___closed__1; x_268 = 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__2_spec__2___closed__2; x_269 = 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__2_spec__2___closed__3; -x_270 = l_Lean_addMacroScope(x_47, x_269, x_46); -x_271 = l_Lean_Name_mkStr2(x_63, x_267); +x_270 = l_Lean_addMacroScope(x_46, x_269, x_51); +x_271 = l_Lean_Name_mkStr2(x_53, x_267); lean_inc(x_271); x_272 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_54); +lean_ctor_set(x_272, 1, x_38); x_273 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_273, 0, x_271); x_274 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_62); +lean_ctor_set(x_274, 1, x_49); x_275 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_275, 0, x_272); lean_ctor_set(x_275, 1, x_274); @@ -8766,17 +8766,17 @@ lean_ctor_set(x_276, 0, x_180); lean_ctor_set(x_276, 1, x_268); lean_ctor_set(x_276, 2, x_270); lean_ctor_set(x_276, 3, x_275); -lean_inc(x_40); +lean_inc(x_54); lean_inc(x_180); -x_277 = l_Lean_Syntax_node2(x_180, x_40, x_266, x_276); +x_277 = l_Lean_Syntax_node2(x_180, x_54, x_266, x_276); lean_inc(x_180); x_278 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_278, 0, x_180); -lean_ctor_set(x_278, 1, x_61); +lean_ctor_set(x_278, 1, x_52); lean_inc(x_180); -x_279 = l_Lean_Syntax_node6(x_180, x_57, x_237, x_238, x_264, x_265, x_277, x_278); +x_279 = l_Lean_Syntax_node6(x_180, x_31, x_237, x_238, x_264, x_265, x_277, x_278); lean_inc(x_180); -x_280 = l_Lean_Syntax_node1(x_180, x_40, x_279); +x_280 = l_Lean_Syntax_node1(x_180, x_54, x_279); x_281 = l_Lean_Syntax_node4(x_180, x_230, x_19, x_232, x_236, x_280); if (lean_is_scalar(x_20)) { x_282 = lean_alloc_ctor(0, 2, 0); @@ -8791,8 +8791,8 @@ x_285 = lean_nat_dec_lt(x_283, x_284); if (x_285 == 0) { lean_dec(x_284); -lean_dec(x_49); -lean_dec(x_44); +lean_dec(x_56); +lean_dec(x_40); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_282; @@ -8806,8 +8806,8 @@ x_286 = lean_nat_dec_le(x_284, x_284); if (x_286 == 0) { lean_dec(x_284); -lean_dec(x_49); -lean_dec(x_44); +lean_dec(x_56); +lean_dec(x_40); lean_dec(x_24); lean_dec_ref(x_23); x_11 = x_282; @@ -8823,7 +8823,7 @@ lean_dec(x_284); lean_inc_ref(x_9); lean_inc(x_3); lean_inc(x_1); -x_289 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_44, x_49, x_3, x_4, x_23, x_287, x_288, x_282, x_9, x_181); +x_289 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__2(x_17, x_1, x_24, x_56, x_40, x_3, x_4, x_23, x_287, x_288, x_282, x_9, x_181); lean_dec_ref(x_23); x_290 = lean_ctor_get(x_289, 0); lean_inc(x_290); @@ -8840,7 +8840,7 @@ goto block_16; block_335: { lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; -x_327 = 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__2_spec__2___lam__1(x_17, x_318, x_9, x_10); +x_327 = 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__2_spec__2___lam__1(x_17, x_323, x_9, x_10); x_328 = lean_ctor_get(x_327, 0); lean_inc(x_328); x_329 = lean_ctor_get(x_327, 1); @@ -8848,53 +8848,53 @@ lean_inc(x_329); lean_dec_ref(x_327); x_330 = l_Lean_mkIdentFrom(x_22, x_326, x_17); lean_dec(x_22); -lean_inc_ref(x_301); -lean_inc(x_306); +lean_inc_ref(x_312); +lean_inc(x_318); lean_inc(x_328); x_331 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_331, 0, x_328); -lean_ctor_set(x_331, 1, x_306); -lean_ctor_set(x_331, 2, x_301); +lean_ctor_set(x_331, 1, x_318); +lean_ctor_set(x_331, 2, x_312); if (lean_obj_tag(x_3) == 0) { lean_object* x_332; x_332 = 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__2_spec__2___closed__53; x_27 = x_293; -x_28 = x_297; -x_29 = x_296; -x_30 = x_295; -x_31 = x_294; -x_32 = x_298; -x_33 = x_299; -x_34 = x_300; -x_35 = x_301; -x_36 = x_302; -x_37 = x_303; -x_38 = x_304; -x_39 = x_305; -x_40 = x_306; -x_41 = x_330; -x_42 = x_307; -x_43 = x_308; -x_44 = x_309; -x_45 = x_310; -x_46 = x_311; -x_47 = x_312; -x_48 = x_313; -x_49 = x_314; -x_50 = x_315; -x_51 = x_331; -x_52 = x_329; -x_53 = x_317; -x_54 = x_316; -x_55 = x_318; +x_28 = x_294; +x_29 = x_295; +x_30 = x_329; +x_31 = x_296; +x_32 = x_297; +x_33 = x_298; +x_34 = x_328; +x_35 = x_299; +x_36 = x_301; +x_37 = x_300; +x_38 = x_302; +x_39 = x_303; +x_40 = x_304; +x_41 = x_305; +x_42 = x_306; +x_43 = x_307; +x_44 = x_308; +x_45 = x_309; +x_46 = x_310; +x_47 = x_311; +x_48 = x_312; +x_49 = x_313; +x_50 = x_331; +x_51 = x_314; +x_52 = x_315; +x_53 = x_316; +x_54 = x_318; +x_55 = x_317; x_56 = x_319; -x_57 = x_322; -x_58 = x_321; -x_59 = x_320; -x_60 = x_328; -x_61 = x_324; -x_62 = x_323; +x_57 = x_330; +x_58 = x_320; +x_59 = x_322; +x_60 = x_321; +x_61 = x_323; +x_62 = x_324; x_63 = x_325; x_64 = x_332; goto block_292; @@ -8906,41 +8906,41 @@ x_333 = lean_ctor_get(x_3, 0); lean_inc(x_333); x_334 = l_Array_mkArray1___redArg(x_333); x_27 = x_293; -x_28 = x_297; -x_29 = x_296; -x_30 = x_295; -x_31 = x_294; -x_32 = x_298; -x_33 = x_299; -x_34 = x_300; -x_35 = x_301; -x_36 = x_302; -x_37 = x_303; -x_38 = x_304; -x_39 = x_305; -x_40 = x_306; -x_41 = x_330; -x_42 = x_307; -x_43 = x_308; -x_44 = x_309; -x_45 = x_310; -x_46 = x_311; -x_47 = x_312; -x_48 = x_313; -x_49 = x_314; -x_50 = x_315; -x_51 = x_331; -x_52 = x_329; -x_53 = x_317; -x_54 = x_316; -x_55 = x_318; +x_28 = x_294; +x_29 = x_295; +x_30 = x_329; +x_31 = x_296; +x_32 = x_297; +x_33 = x_298; +x_34 = x_328; +x_35 = x_299; +x_36 = x_301; +x_37 = x_300; +x_38 = x_302; +x_39 = x_303; +x_40 = x_304; +x_41 = x_305; +x_42 = x_306; +x_43 = x_307; +x_44 = x_308; +x_45 = x_309; +x_46 = x_310; +x_47 = x_311; +x_48 = x_312; +x_49 = x_313; +x_50 = x_331; +x_51 = x_314; +x_52 = x_315; +x_53 = x_316; +x_54 = x_318; +x_55 = x_317; x_56 = x_319; -x_57 = x_322; -x_58 = x_321; -x_59 = x_320; -x_60 = x_328; -x_61 = x_324; -x_62 = x_323; +x_57 = x_330; +x_58 = x_320; +x_59 = x_322; +x_60 = x_321; +x_61 = x_323; +x_62 = x_324; x_63 = x_325; x_64 = x_334; goto block_292; @@ -8949,21 +8949,21 @@ goto block_292; block_812: { lean_object* x_374; lean_object* x_375; lean_object* x_376; uint8_t x_377; -x_374 = l_Lean_replaceRef(x_19, x_364); +x_374 = l_Lean_replaceRef(x_19, x_370); lean_inc(x_374); lean_inc(x_340); -lean_inc(x_352); -lean_inc(x_356); -lean_inc(x_357); -lean_inc(x_358); +lean_inc(x_336); +lean_inc(x_361); +lean_inc(x_355); +lean_inc(x_337); x_375 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_375, 0, x_358); -lean_ctor_set(x_375, 1, x_357); -lean_ctor_set(x_375, 2, x_356); -lean_ctor_set(x_375, 3, x_352); +lean_ctor_set(x_375, 0, x_337); +lean_ctor_set(x_375, 1, x_355); +lean_ctor_set(x_375, 2, x_361); +lean_ctor_set(x_375, 3, x_336); lean_ctor_set(x_375, 4, x_340); lean_ctor_set(x_375, 5, x_374); -x_376 = 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__2_spec__2___lam__1(x_17, x_374, x_375, x_360); +x_376 = 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__2_spec__2___lam__1(x_17, x_374, x_375, x_347); lean_dec_ref(x_375); lean_dec(x_374); x_377 = !lean_is_exclusive(x_376); @@ -8973,36 +8973,36 @@ lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; x_378 = lean_ctor_get(x_376, 0); x_379 = lean_ctor_get(x_376, 1); x_380 = 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__2_spec__2___closed__25; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_381 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_380); +lean_inc_ref(x_367); +x_381 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_380); x_382 = 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__2_spec__2___closed__26; lean_inc(x_378); lean_ctor_set_tag(x_376, 2); lean_ctor_set(x_376, 1, x_382); x_383 = 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__4_spec__4___closed__1; x_384 = 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__4_spec__4___closed__2; +lean_inc(x_361); +lean_inc(x_355); +x_385 = l_Lean_addMacroScope(x_355, x_384, x_361); lean_inc(x_356); -lean_inc(x_357); -x_385 = l_Lean_addMacroScope(x_357, x_384, x_356); -lean_inc(x_371); lean_inc(x_378); x_386 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_386, 0, x_378); lean_ctor_set(x_386, 1, x_383); lean_ctor_set(x_386, 2, x_385); -lean_ctor_set(x_386, 3, x_371); +lean_ctor_set(x_386, 3, x_356); x_387 = 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__4_spec__4___closed__3; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_388 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_387); +lean_inc_ref(x_367); +x_388 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_387); x_389 = 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__4_spec__4___closed__4; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_390 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_389); +lean_inc_ref(x_367); +x_390 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_389); x_391 = 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__4_spec__4___closed__5; lean_inc(x_378); x_392 = lean_alloc_ctor(2, 2, 0); @@ -9011,37 +9011,37 @@ lean_ctor_set(x_392, 1, x_391); x_393 = 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__4_spec__4___closed__7; x_394 = 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__4_spec__4___closed__9; x_395 = lean_box(0); -lean_inc(x_356); -lean_inc(x_357); -x_396 = l_Lean_addMacroScope(x_357, x_395, x_356); -lean_inc_ref(x_372); -x_397 = l_Lean_Name_mkStr1(x_372); +lean_inc(x_361); +lean_inc(x_355); +x_396 = l_Lean_addMacroScope(x_355, x_395, x_361); +lean_inc_ref(x_363); +x_397 = l_Lean_Name_mkStr1(x_363); x_398 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_398, 0, x_397); -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_399 = l_Lean_Name_mkStr3(x_362, x_342, x_341); +lean_inc_ref(x_367); +x_399 = l_Lean_Name_mkStr3(x_367, x_342, x_364); x_400 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_400, 0, x_399); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_401 = l_Lean_Name_mkStr2(x_362, x_342); +lean_inc_ref(x_367); +x_401 = l_Lean_Name_mkStr2(x_367, x_342); x_402 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_402, 0, x_401); x_403 = 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__4_spec__4___closed__10; -lean_inc_ref(x_362); -x_404 = l_Lean_Name_mkStr2(x_362, x_403); +lean_inc_ref(x_367); +x_404 = l_Lean_Name_mkStr2(x_367, x_403); x_405 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_405, 0, x_404); -lean_inc_ref(x_362); -x_406 = l_Lean_Name_mkStr1(x_362); +lean_inc_ref(x_367); +x_406 = l_Lean_Name_mkStr1(x_367); x_407 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_407, 0, x_406); -lean_inc(x_371); +lean_inc(x_356); x_408 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_408, 0, x_407); -lean_ctor_set(x_408, 1, x_371); +lean_ctor_set(x_408, 1, x_356); x_409 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_409, 0, x_405); lean_ctor_set(x_409, 1, x_408); @@ -9068,19 +9068,19 @@ x_416 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__ x_417 = 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__4_spec__4___closed__13; x_418 = 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__4_spec__4___closed__14; x_419 = 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__4_spec__4___closed__15; -lean_inc(x_356); -lean_inc(x_357); -x_420 = l_Lean_addMacroScope(x_357, x_419, x_356); -lean_inc_ref(x_372); -x_421 = l_Lean_Name_mkStr3(x_372, x_417, x_418); -lean_inc(x_363); +lean_inc(x_361); +lean_inc(x_355); +x_420 = l_Lean_addMacroScope(x_355, x_419, x_361); +lean_inc_ref(x_363); +x_421 = l_Lean_Name_mkStr3(x_363, x_417, x_418); +lean_inc(x_345); x_422 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_422, 0, x_421); -lean_ctor_set(x_422, 1, x_363); -lean_inc(x_371); +lean_ctor_set(x_422, 1, x_345); +lean_inc(x_356); x_423 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_423, 0, x_422); -lean_ctor_set(x_423, 1, x_371); +lean_ctor_set(x_423, 1, x_356); lean_inc(x_378); x_424 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_424, 0, x_378); @@ -9088,12 +9088,12 @@ lean_ctor_set(x_424, 1, x_416); lean_ctor_set(x_424, 2, x_420); lean_ctor_set(x_424, 3, x_423); lean_inc(x_24); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_378); -x_425 = l_Lean_Syntax_node1(x_378, x_351, x_24); -lean_inc(x_353); +x_425 = l_Lean_Syntax_node1(x_378, x_365, x_24); +lean_inc(x_349); lean_inc(x_378); -x_426 = l_Lean_Syntax_node2(x_378, x_353, x_424, x_425); +x_426 = l_Lean_Syntax_node2(x_378, x_349, x_424, x_425); x_427 = 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__4_spec__4___closed__16; lean_inc(x_378); x_428 = lean_alloc_ctor(2, 2, 0); @@ -9101,21 +9101,21 @@ lean_ctor_set(x_428, 0, x_378); lean_ctor_set(x_428, 1, x_427); lean_inc(x_378); x_429 = l_Lean_Syntax_node3(x_378, x_388, x_415, x_426, x_428); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_378); -x_430 = l_Lean_Syntax_node1(x_378, x_351, x_429); +x_430 = l_Lean_Syntax_node1(x_378, x_365, x_429); lean_inc(x_381); x_431 = l_Lean_Syntax_node4(x_378, x_381, x_19, x_376, x_386, x_430); -x_432 = l_Lean_replaceRef(x_431, x_364); -lean_dec(x_364); +x_432 = l_Lean_replaceRef(x_431, x_370); +lean_dec(x_370); lean_inc(x_432); -lean_inc(x_356); -lean_inc(x_357); +lean_inc(x_361); +lean_inc(x_355); x_433 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_433, 0, x_358); -lean_ctor_set(x_433, 1, x_357); -lean_ctor_set(x_433, 2, x_356); -lean_ctor_set(x_433, 3, x_352); +lean_ctor_set(x_433, 0, x_337); +lean_ctor_set(x_433, 1, x_355); +lean_ctor_set(x_433, 2, x_361); +lean_ctor_set(x_433, 3, x_336); lean_ctor_set(x_433, 4, x_340); lean_ctor_set(x_433, 5, x_432); x_434 = 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__2_spec__2___lam__1(x_17, x_432, x_433, x_379); @@ -9127,238 +9127,238 @@ if (x_435 == 0) lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; 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_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; x_436 = lean_ctor_get(x_434, 0); x_437 = lean_ctor_get(x_434, 1); -lean_inc_ref(x_345); -x_438 = l_Array_append___redArg(x_345, x_373); +lean_inc_ref(x_357); +x_438 = l_Array_append___redArg(x_357, x_373); lean_dec_ref(x_373); -lean_inc(x_351); -lean_inc(x_344); +lean_inc(x_365); +lean_inc(x_359); x_439 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_439, 0, x_344); -lean_ctor_set(x_439, 1, x_351); +lean_ctor_set(x_439, 0, x_359); +lean_ctor_set(x_439, 1, x_365); lean_ctor_set(x_439, 2, x_438); -lean_inc_n(x_365, 6); -lean_inc(x_344); -x_440 = l_Lean_Syntax_node7(x_344, x_346, x_365, x_365, x_439, x_365, x_365, x_365, x_365); +lean_inc_n(x_358, 6); +lean_inc(x_359); +x_440 = l_Lean_Syntax_node7(x_359, x_338, x_358, x_358, x_439, x_358, x_358, x_358, x_358); x_441 = 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__2_spec__2___closed__5; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_442 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_441); +lean_inc_ref(x_367); +x_442 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_441); x_443 = 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__2_spec__2___closed__6; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_444 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_443); -lean_inc(x_365); -lean_inc(x_344); -x_445 = l_Lean_Syntax_node1(x_344, x_444, x_365); -lean_inc(x_344); +lean_inc_ref(x_367); +x_444 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_443); +lean_inc(x_358); +lean_inc(x_359); +x_445 = l_Lean_Syntax_node1(x_359, x_444, x_358); +lean_inc(x_359); lean_ctor_set_tag(x_434, 2); lean_ctor_set(x_434, 1, x_441); -lean_ctor_set(x_434, 0, x_344); +lean_ctor_set(x_434, 0, x_359); +lean_inc(x_358); +lean_inc(x_359); +x_446 = l_Lean_Syntax_node2(x_359, x_352, x_360, x_358); lean_inc(x_365); -lean_inc(x_344); -x_446 = l_Lean_Syntax_node2(x_344, x_343, x_350, x_365); -lean_inc(x_351); -lean_inc(x_344); -x_447 = l_Lean_Syntax_node1(x_344, x_351, x_446); +lean_inc(x_359); +x_447 = l_Lean_Syntax_node1(x_359, x_365, x_446); x_448 = l_Lake_configField___closed__19; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_449 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_448); -lean_inc_ref(x_355); -lean_inc(x_344); +lean_inc_ref(x_367); +x_449 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_448); +lean_inc_ref(x_353); +lean_inc(x_359); x_450 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_450, 0, x_344); -lean_ctor_set(x_450, 1, x_355); +lean_ctor_set(x_450, 0, x_359); +lean_ctor_set(x_450, 1, x_353); x_451 = 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__4_spec__4___closed__17; x_452 = 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__4_spec__4___closed__18; x_453 = 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__4_spec__4___closed__19; -lean_inc(x_356); -lean_inc(x_357); -x_454 = l_Lean_addMacroScope(x_357, x_453, x_356); -lean_inc_ref(x_372); -x_455 = l_Lean_Name_mkStr2(x_372, x_451); -lean_inc(x_363); +lean_inc(x_361); +lean_inc(x_355); +x_454 = l_Lean_addMacroScope(x_355, x_453, x_361); +lean_inc_ref(x_363); +x_455 = l_Lean_Name_mkStr2(x_363, x_451); +lean_inc(x_345); lean_inc(x_455); x_456 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_456, 0, x_455); -lean_ctor_set(x_456, 1, x_363); +lean_ctor_set(x_456, 1, x_345); x_457 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_457, 0, x_455); -lean_inc(x_371); +lean_inc(x_356); x_458 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_458, 0, x_457); -lean_ctor_set(x_458, 1, x_371); +lean_ctor_set(x_458, 1, x_356); x_459 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_459, 0, x_456); lean_ctor_set(x_459, 1, x_458); -lean_inc(x_344); +lean_inc(x_359); x_460 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_460, 0, x_344); +lean_ctor_set(x_460, 0, x_359); lean_ctor_set(x_460, 1, x_452); lean_ctor_set(x_460, 2, x_454); lean_ctor_set(x_460, 3, x_459); lean_inc(x_1); -lean_inc(x_351); -lean_inc(x_344); -x_461 = l_Lean_Syntax_node2(x_344, x_351, x_1, x_24); -lean_inc(x_344); -x_462 = l_Lean_Syntax_node2(x_344, x_353, x_460, x_461); -lean_inc(x_344); -x_463 = l_Lean_Syntax_node2(x_344, x_369, x_450, x_462); lean_inc(x_365); -lean_inc(x_344); -x_464 = l_Lean_Syntax_node2(x_344, x_449, x_365, x_463); +lean_inc(x_359); +x_461 = l_Lean_Syntax_node2(x_359, x_365, x_1, x_24); +lean_inc(x_359); +x_462 = l_Lean_Syntax_node2(x_359, x_349, x_460, x_461); +lean_inc(x_359); +x_463 = l_Lean_Syntax_node2(x_359, x_371, x_450, x_462); +lean_inc(x_358); +lean_inc(x_359); +x_464 = l_Lean_Syntax_node2(x_359, x_449, x_358, x_463); x_465 = 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__2_spec__2___closed__18; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_466 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_465); -lean_inc_ref(x_366); -lean_inc(x_344); +lean_inc_ref(x_367); +x_466 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_465); +lean_inc_ref(x_369); +lean_inc(x_359); x_467 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_467, 0, x_344); -lean_ctor_set(x_467, 1, x_366); +lean_ctor_set(x_467, 0, x_359); +lean_ctor_set(x_467, 1, x_369); x_468 = 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__2_spec__2___closed__20; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_469 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_468); +lean_inc_ref(x_367); +x_469 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_468); x_470 = 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__2_spec__2___closed__21; -lean_inc(x_344); +lean_inc(x_359); x_471 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_471, 0, x_344); +lean_ctor_set(x_471, 0, x_359); lean_ctor_set(x_471, 1, x_470); -lean_inc(x_351); -lean_inc(x_344); -x_472 = l_Lean_Syntax_node1(x_344, x_351, x_354); +lean_inc(x_365); +lean_inc(x_359); +x_472 = l_Lean_Syntax_node1(x_359, x_365, x_366); x_473 = 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__2_spec__2___closed__22; -lean_inc(x_344); +lean_inc(x_359); x_474 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_474, 0, x_344); +lean_ctor_set(x_474, 0, x_359); lean_ctor_set(x_474, 1, x_473); -lean_inc(x_344); -x_475 = l_Lean_Syntax_node3(x_344, x_469, x_471, x_472, x_474); +lean_inc(x_359); +x_475 = l_Lean_Syntax_node3(x_359, x_469, x_471, x_472, x_474); x_476 = 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__2_spec__2___closed__23; x_477 = 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__2_spec__2___closed__24; -x_478 = l_Lean_Name_mkStr4(x_362, x_342, x_476, x_477); -lean_inc_n(x_365, 2); -lean_inc(x_344); -x_479 = l_Lean_Syntax_node2(x_344, x_478, x_365, x_365); -lean_inc(x_365); -lean_inc(x_344); -x_480 = l_Lean_Syntax_node4(x_344, x_466, x_467, x_475, x_479, x_365); -lean_inc(x_344); -x_481 = l_Lean_Syntax_node6(x_344, x_442, x_445, x_434, x_365, x_447, x_464, x_480); -x_482 = l_Lean_Syntax_node2(x_344, x_349, x_440, x_481); -x_483 = lean_array_push(x_339, x_482); +x_478 = l_Lean_Name_mkStr4(x_367, x_342, x_476, x_477); +lean_inc_n(x_358, 2); +lean_inc(x_359); +x_479 = l_Lean_Syntax_node2(x_359, x_478, x_358, x_358); +lean_inc(x_358); +lean_inc(x_359); +x_480 = l_Lean_Syntax_node4(x_359, x_466, x_467, x_475, x_479, x_358); +lean_inc(x_359); +x_481 = l_Lean_Syntax_node6(x_359, x_442, x_445, x_434, x_358, x_447, x_464, x_480); +x_482 = l_Lean_Syntax_node2(x_359, x_368, x_440, x_481); +x_483 = lean_array_push(x_344, x_482); lean_inc(x_436); x_484 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_484, 0, x_436); lean_ctor_set(x_484, 1, x_382); x_485 = 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__2_spec__2___closed__28; x_486 = 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__2_spec__2___closed__29; +lean_inc(x_361); +lean_inc(x_355); +x_487 = l_Lean_addMacroScope(x_355, x_486, x_361); lean_inc(x_356); -lean_inc(x_357); -x_487 = l_Lean_addMacroScope(x_357, x_486, x_356); -lean_inc(x_371); lean_inc(x_436); x_488 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_488, 0, x_436); lean_ctor_set(x_488, 1, x_485); lean_ctor_set(x_488, 2, x_487); -lean_ctor_set(x_488, 3, x_371); +lean_ctor_set(x_488, 3, x_356); lean_inc(x_436); x_489 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_489, 0, x_436); -lean_ctor_set(x_489, 1, x_368); -lean_inc(x_351); +lean_ctor_set(x_489, 1, x_343); +lean_inc(x_365); lean_inc(x_436); x_490 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_490, 0, x_436); -lean_ctor_set(x_490, 1, x_351); -lean_ctor_set(x_490, 2, x_345); +lean_ctor_set(x_490, 1, x_365); +lean_ctor_set(x_490, 2, x_357); x_491 = 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__2_spec__2___closed__34; x_492 = 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__2_spec__2___closed__35; +lean_inc(x_361); +lean_inc(x_355); +x_493 = l_Lean_addMacroScope(x_355, x_492, x_361); lean_inc(x_356); -lean_inc(x_357); -x_493 = l_Lean_addMacroScope(x_357, x_492, x_356); -lean_inc(x_371); lean_inc(x_436); x_494 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_494, 0, x_436); lean_ctor_set(x_494, 1, x_491); lean_ctor_set(x_494, 2, x_493); -lean_ctor_set(x_494, 3, x_371); +lean_ctor_set(x_494, 3, x_356); lean_inc_ref(x_490); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_436); -x_495 = l_Lean_Syntax_node2(x_436, x_347, x_494, x_490); +x_495 = l_Lean_Syntax_node2(x_436, x_351, x_494, x_490); lean_inc(x_436); x_496 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_496, 0, x_436); -lean_ctor_set(x_496, 1, x_366); +lean_ctor_set(x_496, 1, x_369); lean_inc_ref(x_490); lean_inc_ref(x_496); -lean_inc(x_336); +lean_inc(x_372); lean_inc(x_436); -x_497 = l_Lean_Syntax_node3(x_436, x_336, x_496, x_490, x_359); +x_497 = l_Lean_Syntax_node3(x_436, x_372, x_496, x_490, x_348); lean_inc_ref_n(x_490, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_436); -x_498 = l_Lean_Syntax_node3(x_436, x_351, x_490, x_490, x_497); +x_498 = l_Lean_Syntax_node3(x_436, x_365, x_490, x_490, x_497); lean_inc(x_498); -lean_inc(x_338); +lean_inc(x_341); lean_inc(x_436); -x_499 = l_Lean_Syntax_node2(x_436, x_338, x_495, x_498); +x_499 = l_Lean_Syntax_node2(x_436, x_341, x_495, x_498); x_500 = 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__2_spec__2___closed__38; x_501 = 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__2_spec__2___closed__39; +lean_inc(x_361); +lean_inc(x_355); +x_502 = l_Lean_addMacroScope(x_355, x_501, x_361); lean_inc(x_356); -lean_inc(x_357); -x_502 = l_Lean_addMacroScope(x_357, x_501, x_356); -lean_inc(x_371); lean_inc(x_436); x_503 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_503, 0, x_436); lean_ctor_set(x_503, 1, x_500); lean_ctor_set(x_503, 2, x_502); -lean_ctor_set(x_503, 3, x_371); +lean_ctor_set(x_503, 3, x_356); lean_inc_ref(x_490); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_436); -x_504 = l_Lean_Syntax_node2(x_436, x_347, x_503, x_490); -lean_inc(x_338); +x_504 = l_Lean_Syntax_node2(x_436, x_351, x_503, x_490); +lean_inc(x_341); lean_inc(x_436); -x_505 = l_Lean_Syntax_node2(x_436, x_338, x_504, x_498); +x_505 = l_Lean_Syntax_node2(x_436, x_341, x_504, x_498); x_506 = 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__4_spec__4___closed__21; x_507 = 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__4_spec__4___closed__22; +lean_inc(x_361); +lean_inc(x_355); +x_508 = l_Lean_addMacroScope(x_355, x_507, x_361); lean_inc(x_356); -lean_inc(x_357); -x_508 = l_Lean_addMacroScope(x_357, x_507, x_356); -lean_inc(x_371); lean_inc(x_436); x_509 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_509, 0, x_436); lean_ctor_set(x_509, 1, x_506); lean_ctor_set(x_509, 2, x_508); -lean_ctor_set(x_509, 3, x_371); +lean_ctor_set(x_509, 3, x_356); lean_inc_ref(x_490); lean_inc(x_436); -x_510 = l_Lean_Syntax_node2(x_436, x_347, x_509, x_490); +x_510 = l_Lean_Syntax_node2(x_436, x_351, x_509, x_490); x_511 = 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__4_spec__4___closed__23; x_512 = 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__4_spec__4___closed__24; -lean_inc(x_356); -lean_inc(x_357); -x_513 = l_Lean_addMacroScope(x_357, x_512, x_356); +lean_inc(x_361); +lean_inc(x_355); +x_513 = l_Lean_addMacroScope(x_355, x_512, x_361); x_514 = 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__2_spec__2___closed__48; -lean_inc(x_363); +lean_inc(x_345); x_515 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_515, 0, x_514); -lean_ctor_set(x_515, 1, x_363); -lean_inc(x_371); +lean_ctor_set(x_515, 1, x_345); +lean_inc(x_356); x_516 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_516, 0, x_515); -lean_ctor_set(x_516, 1, x_371); +lean_ctor_set(x_516, 1, x_356); lean_inc(x_436); x_517 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_517, 0, x_436); @@ -9367,40 +9367,40 @@ lean_ctor_set(x_517, 2, x_513); lean_ctor_set(x_517, 3, x_516); lean_inc_ref(x_490); lean_inc(x_436); -x_518 = l_Lean_Syntax_node3(x_436, x_336, x_496, x_490, x_517); +x_518 = l_Lean_Syntax_node3(x_436, x_372, x_496, x_490, x_517); lean_inc_ref_n(x_490, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_436); -x_519 = l_Lean_Syntax_node3(x_436, x_351, x_490, x_490, x_518); +x_519 = l_Lean_Syntax_node3(x_436, x_365, x_490, x_490, x_518); lean_inc(x_436); -x_520 = l_Lean_Syntax_node2(x_436, x_338, x_510, x_519); +x_520 = l_Lean_Syntax_node2(x_436, x_341, x_510, x_519); lean_inc_ref_n(x_490, 3); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_436); -x_521 = l_Lean_Syntax_node6(x_436, x_351, x_499, x_490, x_505, x_490, x_520, x_490); +x_521 = l_Lean_Syntax_node6(x_436, x_365, x_499, x_490, x_505, x_490, x_520, x_490); lean_inc(x_436); -x_522 = l_Lean_Syntax_node1(x_436, x_337, x_521); +x_522 = l_Lean_Syntax_node1(x_436, x_350, x_521); lean_inc_ref(x_490); lean_inc(x_436); -x_523 = l_Lean_Syntax_node1(x_436, x_361, x_490); +x_523 = l_Lean_Syntax_node1(x_436, x_346, x_490); lean_inc(x_436); x_524 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_524, 0, x_436); -lean_ctor_set(x_524, 1, x_355); +lean_ctor_set(x_524, 1, x_353); x_525 = 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__2_spec__2___closed__1; x_526 = 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__2_spec__2___closed__2; x_527 = 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__2_spec__2___closed__3; -x_528 = l_Lean_addMacroScope(x_357, x_527, x_356); -x_529 = l_Lean_Name_mkStr2(x_372, x_525); +x_528 = l_Lean_addMacroScope(x_355, x_527, x_361); +x_529 = l_Lean_Name_mkStr2(x_363, x_525); lean_inc(x_529); x_530 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_530, 0, x_529); -lean_ctor_set(x_530, 1, x_363); +lean_ctor_set(x_530, 1, x_345); x_531 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_531, 0, x_529); x_532 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_532, 0, x_531); -lean_ctor_set(x_532, 1, x_371); +lean_ctor_set(x_532, 1, x_356); x_533 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_533, 0, x_530); lean_ctor_set(x_533, 1, x_532); @@ -9410,17 +9410,17 @@ lean_ctor_set(x_534, 0, x_436); lean_ctor_set(x_534, 1, x_526); lean_ctor_set(x_534, 2, x_528); lean_ctor_set(x_534, 3, x_533); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_436); -x_535 = l_Lean_Syntax_node2(x_436, x_351, x_524, x_534); +x_535 = l_Lean_Syntax_node2(x_436, x_365, x_524, x_534); lean_inc(x_436); x_536 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_536, 0, x_436); -lean_ctor_set(x_536, 1, x_370); +lean_ctor_set(x_536, 1, x_362); lean_inc(x_436); -x_537 = l_Lean_Syntax_node6(x_436, x_367, x_489, x_490, x_522, x_523, x_535, x_536); +x_537 = l_Lean_Syntax_node6(x_436, x_339, x_489, x_490, x_522, x_523, x_535, x_536); lean_inc(x_436); -x_538 = l_Lean_Syntax_node1(x_436, x_351, x_537); +x_538 = l_Lean_Syntax_node1(x_436, x_365, x_537); x_539 = l_Lean_Syntax_node4(x_436, x_381, x_431, x_484, x_488, x_538); x_540 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_540, 0, x_483); @@ -9437,238 +9437,238 @@ x_542 = lean_ctor_get(x_434, 1); lean_inc(x_542); lean_inc(x_541); lean_dec(x_434); -lean_inc_ref(x_345); -x_543 = l_Array_append___redArg(x_345, x_373); +lean_inc_ref(x_357); +x_543 = l_Array_append___redArg(x_357, x_373); lean_dec_ref(x_373); -lean_inc(x_351); -lean_inc(x_344); +lean_inc(x_365); +lean_inc(x_359); x_544 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_544, 0, x_344); -lean_ctor_set(x_544, 1, x_351); +lean_ctor_set(x_544, 0, x_359); +lean_ctor_set(x_544, 1, x_365); lean_ctor_set(x_544, 2, x_543); -lean_inc_n(x_365, 6); -lean_inc(x_344); -x_545 = l_Lean_Syntax_node7(x_344, x_346, x_365, x_365, x_544, x_365, x_365, x_365, x_365); +lean_inc_n(x_358, 6); +lean_inc(x_359); +x_545 = l_Lean_Syntax_node7(x_359, x_338, x_358, x_358, x_544, x_358, x_358, x_358, x_358); x_546 = 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__2_spec__2___closed__5; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_547 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_546); +lean_inc_ref(x_367); +x_547 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_546); x_548 = 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__2_spec__2___closed__6; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_549 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_548); -lean_inc(x_365); -lean_inc(x_344); -x_550 = l_Lean_Syntax_node1(x_344, x_549, x_365); -lean_inc(x_344); +lean_inc_ref(x_367); +x_549 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_548); +lean_inc(x_358); +lean_inc(x_359); +x_550 = l_Lean_Syntax_node1(x_359, x_549, x_358); +lean_inc(x_359); x_551 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_551, 0, x_344); +lean_ctor_set(x_551, 0, x_359); lean_ctor_set(x_551, 1, x_546); +lean_inc(x_358); +lean_inc(x_359); +x_552 = l_Lean_Syntax_node2(x_359, x_352, x_360, x_358); lean_inc(x_365); -lean_inc(x_344); -x_552 = l_Lean_Syntax_node2(x_344, x_343, x_350, x_365); -lean_inc(x_351); -lean_inc(x_344); -x_553 = l_Lean_Syntax_node1(x_344, x_351, x_552); +lean_inc(x_359); +x_553 = l_Lean_Syntax_node1(x_359, x_365, x_552); x_554 = l_Lake_configField___closed__19; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_555 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_554); -lean_inc_ref(x_355); -lean_inc(x_344); +lean_inc_ref(x_367); +x_555 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_554); +lean_inc_ref(x_353); +lean_inc(x_359); x_556 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_556, 0, x_344); -lean_ctor_set(x_556, 1, x_355); +lean_ctor_set(x_556, 0, x_359); +lean_ctor_set(x_556, 1, x_353); x_557 = 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__4_spec__4___closed__17; x_558 = 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__4_spec__4___closed__18; x_559 = 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__4_spec__4___closed__19; -lean_inc(x_356); -lean_inc(x_357); -x_560 = l_Lean_addMacroScope(x_357, x_559, x_356); -lean_inc_ref(x_372); -x_561 = l_Lean_Name_mkStr2(x_372, x_557); -lean_inc(x_363); +lean_inc(x_361); +lean_inc(x_355); +x_560 = l_Lean_addMacroScope(x_355, x_559, x_361); +lean_inc_ref(x_363); +x_561 = l_Lean_Name_mkStr2(x_363, x_557); +lean_inc(x_345); lean_inc(x_561); x_562 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_562, 0, x_561); -lean_ctor_set(x_562, 1, x_363); +lean_ctor_set(x_562, 1, x_345); x_563 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_563, 0, x_561); -lean_inc(x_371); +lean_inc(x_356); x_564 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_564, 0, x_563); -lean_ctor_set(x_564, 1, x_371); +lean_ctor_set(x_564, 1, x_356); x_565 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_565, 0, x_562); lean_ctor_set(x_565, 1, x_564); -lean_inc(x_344); +lean_inc(x_359); x_566 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_566, 0, x_344); +lean_ctor_set(x_566, 0, x_359); lean_ctor_set(x_566, 1, x_558); lean_ctor_set(x_566, 2, x_560); lean_ctor_set(x_566, 3, x_565); lean_inc(x_1); -lean_inc(x_351); -lean_inc(x_344); -x_567 = l_Lean_Syntax_node2(x_344, x_351, x_1, x_24); -lean_inc(x_344); -x_568 = l_Lean_Syntax_node2(x_344, x_353, x_566, x_567); -lean_inc(x_344); -x_569 = l_Lean_Syntax_node2(x_344, x_369, x_556, x_568); lean_inc(x_365); -lean_inc(x_344); -x_570 = l_Lean_Syntax_node2(x_344, x_555, x_365, x_569); +lean_inc(x_359); +x_567 = l_Lean_Syntax_node2(x_359, x_365, x_1, x_24); +lean_inc(x_359); +x_568 = l_Lean_Syntax_node2(x_359, x_349, x_566, x_567); +lean_inc(x_359); +x_569 = l_Lean_Syntax_node2(x_359, x_371, x_556, x_568); +lean_inc(x_358); +lean_inc(x_359); +x_570 = l_Lean_Syntax_node2(x_359, x_555, x_358, x_569); x_571 = 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__2_spec__2___closed__18; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_572 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_571); -lean_inc_ref(x_366); -lean_inc(x_344); +lean_inc_ref(x_367); +x_572 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_571); +lean_inc_ref(x_369); +lean_inc(x_359); x_573 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_573, 0, x_344); -lean_ctor_set(x_573, 1, x_366); +lean_ctor_set(x_573, 0, x_359); +lean_ctor_set(x_573, 1, x_369); x_574 = 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__2_spec__2___closed__20; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_575 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_574); +lean_inc_ref(x_367); +x_575 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_574); x_576 = 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__2_spec__2___closed__21; -lean_inc(x_344); +lean_inc(x_359); x_577 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_577, 0, x_344); +lean_ctor_set(x_577, 0, x_359); lean_ctor_set(x_577, 1, x_576); -lean_inc(x_351); -lean_inc(x_344); -x_578 = l_Lean_Syntax_node1(x_344, x_351, x_354); +lean_inc(x_365); +lean_inc(x_359); +x_578 = l_Lean_Syntax_node1(x_359, x_365, x_366); x_579 = 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__2_spec__2___closed__22; -lean_inc(x_344); +lean_inc(x_359); x_580 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_580, 0, x_344); +lean_ctor_set(x_580, 0, x_359); lean_ctor_set(x_580, 1, x_579); -lean_inc(x_344); -x_581 = l_Lean_Syntax_node3(x_344, x_575, x_577, x_578, x_580); +lean_inc(x_359); +x_581 = l_Lean_Syntax_node3(x_359, x_575, x_577, x_578, x_580); x_582 = 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__2_spec__2___closed__23; x_583 = 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__2_spec__2___closed__24; -x_584 = l_Lean_Name_mkStr4(x_362, x_342, x_582, x_583); -lean_inc_n(x_365, 2); -lean_inc(x_344); -x_585 = l_Lean_Syntax_node2(x_344, x_584, x_365, x_365); -lean_inc(x_365); -lean_inc(x_344); -x_586 = l_Lean_Syntax_node4(x_344, x_572, x_573, x_581, x_585, x_365); -lean_inc(x_344); -x_587 = l_Lean_Syntax_node6(x_344, x_547, x_550, x_551, x_365, x_553, x_570, x_586); -x_588 = l_Lean_Syntax_node2(x_344, x_349, x_545, x_587); -x_589 = lean_array_push(x_339, x_588); +x_584 = l_Lean_Name_mkStr4(x_367, x_342, x_582, x_583); +lean_inc_n(x_358, 2); +lean_inc(x_359); +x_585 = l_Lean_Syntax_node2(x_359, x_584, x_358, x_358); +lean_inc(x_358); +lean_inc(x_359); +x_586 = l_Lean_Syntax_node4(x_359, x_572, x_573, x_581, x_585, x_358); +lean_inc(x_359); +x_587 = l_Lean_Syntax_node6(x_359, x_547, x_550, x_551, x_358, x_553, x_570, x_586); +x_588 = l_Lean_Syntax_node2(x_359, x_368, x_545, x_587); +x_589 = lean_array_push(x_344, x_588); lean_inc(x_541); x_590 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_590, 0, x_541); lean_ctor_set(x_590, 1, x_382); x_591 = 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__2_spec__2___closed__28; x_592 = 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__2_spec__2___closed__29; +lean_inc(x_361); +lean_inc(x_355); +x_593 = l_Lean_addMacroScope(x_355, x_592, x_361); lean_inc(x_356); -lean_inc(x_357); -x_593 = l_Lean_addMacroScope(x_357, x_592, x_356); -lean_inc(x_371); lean_inc(x_541); x_594 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_594, 0, x_541); lean_ctor_set(x_594, 1, x_591); lean_ctor_set(x_594, 2, x_593); -lean_ctor_set(x_594, 3, x_371); +lean_ctor_set(x_594, 3, x_356); lean_inc(x_541); x_595 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_595, 0, x_541); -lean_ctor_set(x_595, 1, x_368); -lean_inc(x_351); +lean_ctor_set(x_595, 1, x_343); +lean_inc(x_365); lean_inc(x_541); x_596 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_596, 0, x_541); -lean_ctor_set(x_596, 1, x_351); -lean_ctor_set(x_596, 2, x_345); +lean_ctor_set(x_596, 1, x_365); +lean_ctor_set(x_596, 2, x_357); x_597 = 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__2_spec__2___closed__34; x_598 = 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__2_spec__2___closed__35; +lean_inc(x_361); +lean_inc(x_355); +x_599 = l_Lean_addMacroScope(x_355, x_598, x_361); lean_inc(x_356); -lean_inc(x_357); -x_599 = l_Lean_addMacroScope(x_357, x_598, x_356); -lean_inc(x_371); lean_inc(x_541); x_600 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_600, 0, x_541); lean_ctor_set(x_600, 1, x_597); lean_ctor_set(x_600, 2, x_599); -lean_ctor_set(x_600, 3, x_371); +lean_ctor_set(x_600, 3, x_356); lean_inc_ref(x_596); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_541); -x_601 = l_Lean_Syntax_node2(x_541, x_347, x_600, x_596); +x_601 = l_Lean_Syntax_node2(x_541, x_351, x_600, x_596); lean_inc(x_541); x_602 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_602, 0, x_541); -lean_ctor_set(x_602, 1, x_366); +lean_ctor_set(x_602, 1, x_369); lean_inc_ref(x_596); lean_inc_ref(x_602); -lean_inc(x_336); +lean_inc(x_372); lean_inc(x_541); -x_603 = l_Lean_Syntax_node3(x_541, x_336, x_602, x_596, x_359); +x_603 = l_Lean_Syntax_node3(x_541, x_372, x_602, x_596, x_348); lean_inc_ref_n(x_596, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_541); -x_604 = l_Lean_Syntax_node3(x_541, x_351, x_596, x_596, x_603); +x_604 = l_Lean_Syntax_node3(x_541, x_365, x_596, x_596, x_603); lean_inc(x_604); -lean_inc(x_338); +lean_inc(x_341); lean_inc(x_541); -x_605 = l_Lean_Syntax_node2(x_541, x_338, x_601, x_604); +x_605 = l_Lean_Syntax_node2(x_541, x_341, x_601, x_604); x_606 = 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__2_spec__2___closed__38; x_607 = 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__2_spec__2___closed__39; +lean_inc(x_361); +lean_inc(x_355); +x_608 = l_Lean_addMacroScope(x_355, x_607, x_361); lean_inc(x_356); -lean_inc(x_357); -x_608 = l_Lean_addMacroScope(x_357, x_607, x_356); -lean_inc(x_371); lean_inc(x_541); x_609 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_609, 0, x_541); lean_ctor_set(x_609, 1, x_606); lean_ctor_set(x_609, 2, x_608); -lean_ctor_set(x_609, 3, x_371); +lean_ctor_set(x_609, 3, x_356); lean_inc_ref(x_596); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_541); -x_610 = l_Lean_Syntax_node2(x_541, x_347, x_609, x_596); -lean_inc(x_338); +x_610 = l_Lean_Syntax_node2(x_541, x_351, x_609, x_596); +lean_inc(x_341); lean_inc(x_541); -x_611 = l_Lean_Syntax_node2(x_541, x_338, x_610, x_604); +x_611 = l_Lean_Syntax_node2(x_541, x_341, x_610, x_604); x_612 = 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__4_spec__4___closed__21; x_613 = 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__4_spec__4___closed__22; +lean_inc(x_361); +lean_inc(x_355); +x_614 = l_Lean_addMacroScope(x_355, x_613, x_361); lean_inc(x_356); -lean_inc(x_357); -x_614 = l_Lean_addMacroScope(x_357, x_613, x_356); -lean_inc(x_371); lean_inc(x_541); x_615 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_615, 0, x_541); lean_ctor_set(x_615, 1, x_612); lean_ctor_set(x_615, 2, x_614); -lean_ctor_set(x_615, 3, x_371); +lean_ctor_set(x_615, 3, x_356); lean_inc_ref(x_596); lean_inc(x_541); -x_616 = l_Lean_Syntax_node2(x_541, x_347, x_615, x_596); +x_616 = l_Lean_Syntax_node2(x_541, x_351, x_615, x_596); x_617 = 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__4_spec__4___closed__23; x_618 = 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__4_spec__4___closed__24; -lean_inc(x_356); -lean_inc(x_357); -x_619 = l_Lean_addMacroScope(x_357, x_618, x_356); +lean_inc(x_361); +lean_inc(x_355); +x_619 = l_Lean_addMacroScope(x_355, x_618, x_361); x_620 = 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__2_spec__2___closed__48; -lean_inc(x_363); +lean_inc(x_345); x_621 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_621, 0, x_620); -lean_ctor_set(x_621, 1, x_363); -lean_inc(x_371); +lean_ctor_set(x_621, 1, x_345); +lean_inc(x_356); x_622 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_622, 0, x_621); -lean_ctor_set(x_622, 1, x_371); +lean_ctor_set(x_622, 1, x_356); lean_inc(x_541); x_623 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_623, 0, x_541); @@ -9677,40 +9677,40 @@ lean_ctor_set(x_623, 2, x_619); lean_ctor_set(x_623, 3, x_622); lean_inc_ref(x_596); lean_inc(x_541); -x_624 = l_Lean_Syntax_node3(x_541, x_336, x_602, x_596, x_623); +x_624 = l_Lean_Syntax_node3(x_541, x_372, x_602, x_596, x_623); lean_inc_ref_n(x_596, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_541); -x_625 = l_Lean_Syntax_node3(x_541, x_351, x_596, x_596, x_624); +x_625 = l_Lean_Syntax_node3(x_541, x_365, x_596, x_596, x_624); lean_inc(x_541); -x_626 = l_Lean_Syntax_node2(x_541, x_338, x_616, x_625); +x_626 = l_Lean_Syntax_node2(x_541, x_341, x_616, x_625); lean_inc_ref_n(x_596, 3); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_541); -x_627 = l_Lean_Syntax_node6(x_541, x_351, x_605, x_596, x_611, x_596, x_626, x_596); +x_627 = l_Lean_Syntax_node6(x_541, x_365, x_605, x_596, x_611, x_596, x_626, x_596); lean_inc(x_541); -x_628 = l_Lean_Syntax_node1(x_541, x_337, x_627); +x_628 = l_Lean_Syntax_node1(x_541, x_350, x_627); lean_inc_ref(x_596); lean_inc(x_541); -x_629 = l_Lean_Syntax_node1(x_541, x_361, x_596); +x_629 = l_Lean_Syntax_node1(x_541, x_346, x_596); lean_inc(x_541); x_630 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_630, 0, x_541); -lean_ctor_set(x_630, 1, x_355); +lean_ctor_set(x_630, 1, x_353); x_631 = 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__2_spec__2___closed__1; x_632 = 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__2_spec__2___closed__2; x_633 = 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__2_spec__2___closed__3; -x_634 = l_Lean_addMacroScope(x_357, x_633, x_356); -x_635 = l_Lean_Name_mkStr2(x_372, x_631); +x_634 = l_Lean_addMacroScope(x_355, x_633, x_361); +x_635 = l_Lean_Name_mkStr2(x_363, x_631); lean_inc(x_635); x_636 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_636, 0, x_635); -lean_ctor_set(x_636, 1, x_363); +lean_ctor_set(x_636, 1, x_345); x_637 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_637, 0, x_635); x_638 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_638, 0, x_637); -lean_ctor_set(x_638, 1, x_371); +lean_ctor_set(x_638, 1, x_356); x_639 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_639, 0, x_636); lean_ctor_set(x_639, 1, x_638); @@ -9720,17 +9720,17 @@ lean_ctor_set(x_640, 0, x_541); lean_ctor_set(x_640, 1, x_632); lean_ctor_set(x_640, 2, x_634); lean_ctor_set(x_640, 3, x_639); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_541); -x_641 = l_Lean_Syntax_node2(x_541, x_351, x_630, x_640); +x_641 = l_Lean_Syntax_node2(x_541, x_365, x_630, x_640); lean_inc(x_541); x_642 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_642, 0, x_541); -lean_ctor_set(x_642, 1, x_370); +lean_ctor_set(x_642, 1, x_362); lean_inc(x_541); -x_643 = l_Lean_Syntax_node6(x_541, x_367, x_595, x_596, x_628, x_629, x_641, x_642); +x_643 = l_Lean_Syntax_node6(x_541, x_339, x_595, x_596, x_628, x_629, x_641, x_642); lean_inc(x_541); -x_644 = l_Lean_Syntax_node1(x_541, x_351, x_643); +x_644 = l_Lean_Syntax_node1(x_541, x_365, x_643); x_645 = l_Lean_Syntax_node4(x_541, x_381, x_431, x_590, x_594, x_644); x_646 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_646, 0, x_589); @@ -9749,10 +9749,10 @@ lean_inc(x_648); lean_inc(x_647); lean_dec(x_376); x_649 = 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__2_spec__2___closed__25; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_650 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_649); +lean_inc_ref(x_367); +x_650 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_649); x_651 = 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__2_spec__2___closed__26; lean_inc(x_647); x_652 = lean_alloc_ctor(2, 2, 0); @@ -9760,26 +9760,26 @@ lean_ctor_set(x_652, 0, x_647); lean_ctor_set(x_652, 1, x_651); x_653 = 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__4_spec__4___closed__1; x_654 = 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__4_spec__4___closed__2; +lean_inc(x_361); +lean_inc(x_355); +x_655 = l_Lean_addMacroScope(x_355, x_654, x_361); lean_inc(x_356); -lean_inc(x_357); -x_655 = l_Lean_addMacroScope(x_357, x_654, x_356); -lean_inc(x_371); lean_inc(x_647); x_656 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_656, 0, x_647); lean_ctor_set(x_656, 1, x_653); lean_ctor_set(x_656, 2, x_655); -lean_ctor_set(x_656, 3, x_371); +lean_ctor_set(x_656, 3, x_356); x_657 = 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__4_spec__4___closed__3; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_658 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_657); +lean_inc_ref(x_367); +x_658 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_657); x_659 = 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__4_spec__4___closed__4; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_660 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_659); +lean_inc_ref(x_367); +x_660 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_659); x_661 = 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__4_spec__4___closed__5; lean_inc(x_647); x_662 = lean_alloc_ctor(2, 2, 0); @@ -9788,37 +9788,37 @@ lean_ctor_set(x_662, 1, x_661); x_663 = 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__4_spec__4___closed__7; x_664 = 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__4_spec__4___closed__9; x_665 = lean_box(0); -lean_inc(x_356); -lean_inc(x_357); -x_666 = l_Lean_addMacroScope(x_357, x_665, x_356); -lean_inc_ref(x_372); -x_667 = l_Lean_Name_mkStr1(x_372); +lean_inc(x_361); +lean_inc(x_355); +x_666 = l_Lean_addMacroScope(x_355, x_665, x_361); +lean_inc_ref(x_363); +x_667 = l_Lean_Name_mkStr1(x_363); x_668 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_668, 0, x_667); -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_669 = l_Lean_Name_mkStr3(x_362, x_342, x_341); +lean_inc_ref(x_367); +x_669 = l_Lean_Name_mkStr3(x_367, x_342, x_364); x_670 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_670, 0, x_669); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_671 = l_Lean_Name_mkStr2(x_362, x_342); +lean_inc_ref(x_367); +x_671 = l_Lean_Name_mkStr2(x_367, x_342); x_672 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_672, 0, x_671); x_673 = 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__4_spec__4___closed__10; -lean_inc_ref(x_362); -x_674 = l_Lean_Name_mkStr2(x_362, x_673); +lean_inc_ref(x_367); +x_674 = l_Lean_Name_mkStr2(x_367, x_673); x_675 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_675, 0, x_674); -lean_inc_ref(x_362); -x_676 = l_Lean_Name_mkStr1(x_362); +lean_inc_ref(x_367); +x_676 = l_Lean_Name_mkStr1(x_367); x_677 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_677, 0, x_676); -lean_inc(x_371); +lean_inc(x_356); x_678 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_678, 0, x_677); -lean_ctor_set(x_678, 1, x_371); +lean_ctor_set(x_678, 1, x_356); x_679 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_679, 0, x_675); lean_ctor_set(x_679, 1, x_678); @@ -9845,19 +9845,19 @@ x_686 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__ x_687 = 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__4_spec__4___closed__13; x_688 = 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__4_spec__4___closed__14; x_689 = 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__4_spec__4___closed__15; -lean_inc(x_356); -lean_inc(x_357); -x_690 = l_Lean_addMacroScope(x_357, x_689, x_356); -lean_inc_ref(x_372); -x_691 = l_Lean_Name_mkStr3(x_372, x_687, x_688); -lean_inc(x_363); +lean_inc(x_361); +lean_inc(x_355); +x_690 = l_Lean_addMacroScope(x_355, x_689, x_361); +lean_inc_ref(x_363); +x_691 = l_Lean_Name_mkStr3(x_363, x_687, x_688); +lean_inc(x_345); x_692 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_692, 0, x_691); -lean_ctor_set(x_692, 1, x_363); -lean_inc(x_371); +lean_ctor_set(x_692, 1, x_345); +lean_inc(x_356); x_693 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_693, 0, x_692); -lean_ctor_set(x_693, 1, x_371); +lean_ctor_set(x_693, 1, x_356); lean_inc(x_647); x_694 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_694, 0, x_647); @@ -9865,12 +9865,12 @@ lean_ctor_set(x_694, 1, x_686); lean_ctor_set(x_694, 2, x_690); lean_ctor_set(x_694, 3, x_693); lean_inc(x_24); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_647); -x_695 = l_Lean_Syntax_node1(x_647, x_351, x_24); -lean_inc(x_353); +x_695 = l_Lean_Syntax_node1(x_647, x_365, x_24); +lean_inc(x_349); lean_inc(x_647); -x_696 = l_Lean_Syntax_node2(x_647, x_353, x_694, x_695); +x_696 = l_Lean_Syntax_node2(x_647, x_349, x_694, x_695); x_697 = 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__4_spec__4___closed__16; lean_inc(x_647); x_698 = lean_alloc_ctor(2, 2, 0); @@ -9878,21 +9878,21 @@ lean_ctor_set(x_698, 0, x_647); lean_ctor_set(x_698, 1, x_697); lean_inc(x_647); x_699 = l_Lean_Syntax_node3(x_647, x_658, x_685, x_696, x_698); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_647); -x_700 = l_Lean_Syntax_node1(x_647, x_351, x_699); +x_700 = l_Lean_Syntax_node1(x_647, x_365, x_699); lean_inc(x_650); x_701 = l_Lean_Syntax_node4(x_647, x_650, x_19, x_652, x_656, x_700); -x_702 = l_Lean_replaceRef(x_701, x_364); -lean_dec(x_364); +x_702 = l_Lean_replaceRef(x_701, x_370); +lean_dec(x_370); lean_inc(x_702); -lean_inc(x_356); -lean_inc(x_357); +lean_inc(x_361); +lean_inc(x_355); x_703 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_703, 0, x_358); -lean_ctor_set(x_703, 1, x_357); -lean_ctor_set(x_703, 2, x_356); -lean_ctor_set(x_703, 3, x_352); +lean_ctor_set(x_703, 0, x_337); +lean_ctor_set(x_703, 1, x_355); +lean_ctor_set(x_703, 2, x_361); +lean_ctor_set(x_703, 3, x_336); lean_ctor_set(x_703, 4, x_340); lean_ctor_set(x_703, 5, x_702); x_704 = 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__2_spec__2___lam__1(x_17, x_702, x_703, x_648); @@ -9910,243 +9910,243 @@ if (lean_is_exclusive(x_704)) { lean_dec_ref(x_704); x_707 = lean_box(0); } -lean_inc_ref(x_345); -x_708 = l_Array_append___redArg(x_345, x_373); +lean_inc_ref(x_357); +x_708 = l_Array_append___redArg(x_357, x_373); lean_dec_ref(x_373); -lean_inc(x_351); -lean_inc(x_344); +lean_inc(x_365); +lean_inc(x_359); x_709 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_709, 0, x_344); -lean_ctor_set(x_709, 1, x_351); +lean_ctor_set(x_709, 0, x_359); +lean_ctor_set(x_709, 1, x_365); lean_ctor_set(x_709, 2, x_708); -lean_inc_n(x_365, 6); -lean_inc(x_344); -x_710 = l_Lean_Syntax_node7(x_344, x_346, x_365, x_365, x_709, x_365, x_365, x_365, x_365); +lean_inc_n(x_358, 6); +lean_inc(x_359); +x_710 = l_Lean_Syntax_node7(x_359, x_338, x_358, x_358, x_709, x_358, x_358, x_358, x_358); x_711 = 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__2_spec__2___closed__5; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_712 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_711); +lean_inc_ref(x_367); +x_712 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_711); x_713 = 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__2_spec__2___closed__6; -lean_inc_ref(x_348); +lean_inc_ref(x_354); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_714 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_713); -lean_inc(x_365); -lean_inc(x_344); -x_715 = l_Lean_Syntax_node1(x_344, x_714, x_365); -lean_inc(x_344); +lean_inc_ref(x_367); +x_714 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_713); +lean_inc(x_358); +lean_inc(x_359); +x_715 = l_Lean_Syntax_node1(x_359, x_714, x_358); +lean_inc(x_359); if (lean_is_scalar(x_707)) { x_716 = lean_alloc_ctor(2, 2, 0); } else { x_716 = x_707; lean_ctor_set_tag(x_716, 2); } -lean_ctor_set(x_716, 0, x_344); +lean_ctor_set(x_716, 0, x_359); lean_ctor_set(x_716, 1, x_711); +lean_inc(x_358); +lean_inc(x_359); +x_717 = l_Lean_Syntax_node2(x_359, x_352, x_360, x_358); lean_inc(x_365); -lean_inc(x_344); -x_717 = l_Lean_Syntax_node2(x_344, x_343, x_350, x_365); -lean_inc(x_351); -lean_inc(x_344); -x_718 = l_Lean_Syntax_node1(x_344, x_351, x_717); +lean_inc(x_359); +x_718 = l_Lean_Syntax_node1(x_359, x_365, x_717); x_719 = l_Lake_configField___closed__19; -lean_inc_ref(x_341); +lean_inc_ref(x_364); lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_720 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_719); -lean_inc_ref(x_355); -lean_inc(x_344); +lean_inc_ref(x_367); +x_720 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_719); +lean_inc_ref(x_353); +lean_inc(x_359); x_721 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_721, 0, x_344); -lean_ctor_set(x_721, 1, x_355); +lean_ctor_set(x_721, 0, x_359); +lean_ctor_set(x_721, 1, x_353); x_722 = 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__4_spec__4___closed__17; x_723 = 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__4_spec__4___closed__18; x_724 = 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__4_spec__4___closed__19; -lean_inc(x_356); -lean_inc(x_357); -x_725 = l_Lean_addMacroScope(x_357, x_724, x_356); -lean_inc_ref(x_372); -x_726 = l_Lean_Name_mkStr2(x_372, x_722); -lean_inc(x_363); +lean_inc(x_361); +lean_inc(x_355); +x_725 = l_Lean_addMacroScope(x_355, x_724, x_361); +lean_inc_ref(x_363); +x_726 = l_Lean_Name_mkStr2(x_363, x_722); +lean_inc(x_345); lean_inc(x_726); x_727 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_727, 0, x_726); -lean_ctor_set(x_727, 1, x_363); +lean_ctor_set(x_727, 1, x_345); x_728 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_728, 0, x_726); -lean_inc(x_371); +lean_inc(x_356); x_729 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_729, 0, x_728); -lean_ctor_set(x_729, 1, x_371); +lean_ctor_set(x_729, 1, x_356); x_730 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_730, 0, x_727); lean_ctor_set(x_730, 1, x_729); -lean_inc(x_344); +lean_inc(x_359); x_731 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_731, 0, x_344); +lean_ctor_set(x_731, 0, x_359); lean_ctor_set(x_731, 1, x_723); lean_ctor_set(x_731, 2, x_725); lean_ctor_set(x_731, 3, x_730); lean_inc(x_1); -lean_inc(x_351); -lean_inc(x_344); -x_732 = l_Lean_Syntax_node2(x_344, x_351, x_1, x_24); -lean_inc(x_344); -x_733 = l_Lean_Syntax_node2(x_344, x_353, x_731, x_732); -lean_inc(x_344); -x_734 = l_Lean_Syntax_node2(x_344, x_369, x_721, x_733); lean_inc(x_365); -lean_inc(x_344); -x_735 = l_Lean_Syntax_node2(x_344, x_720, x_365, x_734); +lean_inc(x_359); +x_732 = l_Lean_Syntax_node2(x_359, x_365, x_1, x_24); +lean_inc(x_359); +x_733 = l_Lean_Syntax_node2(x_359, x_349, x_731, x_732); +lean_inc(x_359); +x_734 = l_Lean_Syntax_node2(x_359, x_371, x_721, x_733); +lean_inc(x_358); +lean_inc(x_359); +x_735 = l_Lean_Syntax_node2(x_359, x_720, x_358, x_734); x_736 = 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__2_spec__2___closed__18; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_737 = l_Lean_Name_mkStr4(x_362, x_342, x_341, x_736); -lean_inc_ref(x_366); -lean_inc(x_344); +lean_inc_ref(x_367); +x_737 = l_Lean_Name_mkStr4(x_367, x_342, x_364, x_736); +lean_inc_ref(x_369); +lean_inc(x_359); x_738 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_738, 0, x_344); -lean_ctor_set(x_738, 1, x_366); +lean_ctor_set(x_738, 0, x_359); +lean_ctor_set(x_738, 1, x_369); x_739 = 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__2_spec__2___closed__20; lean_inc_ref(x_342); -lean_inc_ref(x_362); -x_740 = l_Lean_Name_mkStr4(x_362, x_342, x_348, x_739); +lean_inc_ref(x_367); +x_740 = l_Lean_Name_mkStr4(x_367, x_342, x_354, x_739); x_741 = 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__2_spec__2___closed__21; -lean_inc(x_344); +lean_inc(x_359); x_742 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_742, 0, x_344); +lean_ctor_set(x_742, 0, x_359); lean_ctor_set(x_742, 1, x_741); -lean_inc(x_351); -lean_inc(x_344); -x_743 = l_Lean_Syntax_node1(x_344, x_351, x_354); +lean_inc(x_365); +lean_inc(x_359); +x_743 = l_Lean_Syntax_node1(x_359, x_365, x_366); x_744 = 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__2_spec__2___closed__22; -lean_inc(x_344); +lean_inc(x_359); x_745 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_745, 0, x_344); +lean_ctor_set(x_745, 0, x_359); lean_ctor_set(x_745, 1, x_744); -lean_inc(x_344); -x_746 = l_Lean_Syntax_node3(x_344, x_740, x_742, x_743, x_745); +lean_inc(x_359); +x_746 = l_Lean_Syntax_node3(x_359, x_740, x_742, x_743, x_745); x_747 = 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__2_spec__2___closed__23; x_748 = 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__2_spec__2___closed__24; -x_749 = l_Lean_Name_mkStr4(x_362, x_342, x_747, x_748); -lean_inc_n(x_365, 2); -lean_inc(x_344); -x_750 = l_Lean_Syntax_node2(x_344, x_749, x_365, x_365); -lean_inc(x_365); -lean_inc(x_344); -x_751 = l_Lean_Syntax_node4(x_344, x_737, x_738, x_746, x_750, x_365); -lean_inc(x_344); -x_752 = l_Lean_Syntax_node6(x_344, x_712, x_715, x_716, x_365, x_718, x_735, x_751); -x_753 = l_Lean_Syntax_node2(x_344, x_349, x_710, x_752); -x_754 = lean_array_push(x_339, x_753); +x_749 = l_Lean_Name_mkStr4(x_367, x_342, x_747, x_748); +lean_inc_n(x_358, 2); +lean_inc(x_359); +x_750 = l_Lean_Syntax_node2(x_359, x_749, x_358, x_358); +lean_inc(x_358); +lean_inc(x_359); +x_751 = l_Lean_Syntax_node4(x_359, x_737, x_738, x_746, x_750, x_358); +lean_inc(x_359); +x_752 = l_Lean_Syntax_node6(x_359, x_712, x_715, x_716, x_358, x_718, x_735, x_751); +x_753 = l_Lean_Syntax_node2(x_359, x_368, x_710, x_752); +x_754 = lean_array_push(x_344, x_753); lean_inc(x_705); x_755 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_755, 0, x_705); lean_ctor_set(x_755, 1, x_651); x_756 = 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__2_spec__2___closed__28; x_757 = 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__2_spec__2___closed__29; +lean_inc(x_361); +lean_inc(x_355); +x_758 = l_Lean_addMacroScope(x_355, x_757, x_361); lean_inc(x_356); -lean_inc(x_357); -x_758 = l_Lean_addMacroScope(x_357, x_757, x_356); -lean_inc(x_371); lean_inc(x_705); x_759 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_759, 0, x_705); lean_ctor_set(x_759, 1, x_756); lean_ctor_set(x_759, 2, x_758); -lean_ctor_set(x_759, 3, x_371); +lean_ctor_set(x_759, 3, x_356); lean_inc(x_705); x_760 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_760, 0, x_705); -lean_ctor_set(x_760, 1, x_368); -lean_inc(x_351); +lean_ctor_set(x_760, 1, x_343); +lean_inc(x_365); lean_inc(x_705); x_761 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_761, 0, x_705); -lean_ctor_set(x_761, 1, x_351); -lean_ctor_set(x_761, 2, x_345); +lean_ctor_set(x_761, 1, x_365); +lean_ctor_set(x_761, 2, x_357); x_762 = 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__2_spec__2___closed__34; x_763 = 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__2_spec__2___closed__35; +lean_inc(x_361); +lean_inc(x_355); +x_764 = l_Lean_addMacroScope(x_355, x_763, x_361); lean_inc(x_356); -lean_inc(x_357); -x_764 = l_Lean_addMacroScope(x_357, x_763, x_356); -lean_inc(x_371); lean_inc(x_705); x_765 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_765, 0, x_705); lean_ctor_set(x_765, 1, x_762); lean_ctor_set(x_765, 2, x_764); -lean_ctor_set(x_765, 3, x_371); +lean_ctor_set(x_765, 3, x_356); lean_inc_ref(x_761); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_705); -x_766 = l_Lean_Syntax_node2(x_705, x_347, x_765, x_761); +x_766 = l_Lean_Syntax_node2(x_705, x_351, x_765, x_761); lean_inc(x_705); x_767 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_767, 0, x_705); -lean_ctor_set(x_767, 1, x_366); +lean_ctor_set(x_767, 1, x_369); lean_inc_ref(x_761); lean_inc_ref(x_767); -lean_inc(x_336); +lean_inc(x_372); lean_inc(x_705); -x_768 = l_Lean_Syntax_node3(x_705, x_336, x_767, x_761, x_359); +x_768 = l_Lean_Syntax_node3(x_705, x_372, x_767, x_761, x_348); lean_inc_ref_n(x_761, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_705); -x_769 = l_Lean_Syntax_node3(x_705, x_351, x_761, x_761, x_768); +x_769 = l_Lean_Syntax_node3(x_705, x_365, x_761, x_761, x_768); lean_inc(x_769); -lean_inc(x_338); +lean_inc(x_341); lean_inc(x_705); -x_770 = l_Lean_Syntax_node2(x_705, x_338, x_766, x_769); +x_770 = l_Lean_Syntax_node2(x_705, x_341, x_766, x_769); x_771 = 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__2_spec__2___closed__38; x_772 = 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__2_spec__2___closed__39; +lean_inc(x_361); +lean_inc(x_355); +x_773 = l_Lean_addMacroScope(x_355, x_772, x_361); lean_inc(x_356); -lean_inc(x_357); -x_773 = l_Lean_addMacroScope(x_357, x_772, x_356); -lean_inc(x_371); lean_inc(x_705); x_774 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_774, 0, x_705); lean_ctor_set(x_774, 1, x_771); lean_ctor_set(x_774, 2, x_773); -lean_ctor_set(x_774, 3, x_371); +lean_ctor_set(x_774, 3, x_356); lean_inc_ref(x_761); -lean_inc(x_347); +lean_inc(x_351); lean_inc(x_705); -x_775 = l_Lean_Syntax_node2(x_705, x_347, x_774, x_761); -lean_inc(x_338); +x_775 = l_Lean_Syntax_node2(x_705, x_351, x_774, x_761); +lean_inc(x_341); lean_inc(x_705); -x_776 = l_Lean_Syntax_node2(x_705, x_338, x_775, x_769); +x_776 = l_Lean_Syntax_node2(x_705, x_341, x_775, x_769); x_777 = 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__4_spec__4___closed__21; x_778 = 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__4_spec__4___closed__22; +lean_inc(x_361); +lean_inc(x_355); +x_779 = l_Lean_addMacroScope(x_355, x_778, x_361); lean_inc(x_356); -lean_inc(x_357); -x_779 = l_Lean_addMacroScope(x_357, x_778, x_356); -lean_inc(x_371); lean_inc(x_705); x_780 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_780, 0, x_705); lean_ctor_set(x_780, 1, x_777); lean_ctor_set(x_780, 2, x_779); -lean_ctor_set(x_780, 3, x_371); +lean_ctor_set(x_780, 3, x_356); lean_inc_ref(x_761); lean_inc(x_705); -x_781 = l_Lean_Syntax_node2(x_705, x_347, x_780, x_761); +x_781 = l_Lean_Syntax_node2(x_705, x_351, x_780, x_761); x_782 = 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__4_spec__4___closed__23; x_783 = 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__4_spec__4___closed__24; -lean_inc(x_356); -lean_inc(x_357); -x_784 = l_Lean_addMacroScope(x_357, x_783, x_356); +lean_inc(x_361); +lean_inc(x_355); +x_784 = l_Lean_addMacroScope(x_355, x_783, x_361); x_785 = 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__2_spec__2___closed__48; -lean_inc(x_363); +lean_inc(x_345); x_786 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_786, 0, x_785); -lean_ctor_set(x_786, 1, x_363); -lean_inc(x_371); +lean_ctor_set(x_786, 1, x_345); +lean_inc(x_356); x_787 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_787, 0, x_786); -lean_ctor_set(x_787, 1, x_371); +lean_ctor_set(x_787, 1, x_356); lean_inc(x_705); x_788 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_788, 0, x_705); @@ -10155,40 +10155,40 @@ lean_ctor_set(x_788, 2, x_784); lean_ctor_set(x_788, 3, x_787); lean_inc_ref(x_761); lean_inc(x_705); -x_789 = l_Lean_Syntax_node3(x_705, x_336, x_767, x_761, x_788); +x_789 = l_Lean_Syntax_node3(x_705, x_372, x_767, x_761, x_788); lean_inc_ref_n(x_761, 2); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_705); -x_790 = l_Lean_Syntax_node3(x_705, x_351, x_761, x_761, x_789); +x_790 = l_Lean_Syntax_node3(x_705, x_365, x_761, x_761, x_789); lean_inc(x_705); -x_791 = l_Lean_Syntax_node2(x_705, x_338, x_781, x_790); +x_791 = l_Lean_Syntax_node2(x_705, x_341, x_781, x_790); lean_inc_ref_n(x_761, 3); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_705); -x_792 = l_Lean_Syntax_node6(x_705, x_351, x_770, x_761, x_776, x_761, x_791, x_761); +x_792 = l_Lean_Syntax_node6(x_705, x_365, x_770, x_761, x_776, x_761, x_791, x_761); lean_inc(x_705); -x_793 = l_Lean_Syntax_node1(x_705, x_337, x_792); +x_793 = l_Lean_Syntax_node1(x_705, x_350, x_792); lean_inc_ref(x_761); lean_inc(x_705); -x_794 = l_Lean_Syntax_node1(x_705, x_361, x_761); +x_794 = l_Lean_Syntax_node1(x_705, x_346, x_761); lean_inc(x_705); x_795 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_795, 0, x_705); -lean_ctor_set(x_795, 1, x_355); +lean_ctor_set(x_795, 1, x_353); x_796 = 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__2_spec__2___closed__1; x_797 = 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__2_spec__2___closed__2; x_798 = 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__2_spec__2___closed__3; -x_799 = l_Lean_addMacroScope(x_357, x_798, x_356); -x_800 = l_Lean_Name_mkStr2(x_372, x_796); +x_799 = l_Lean_addMacroScope(x_355, x_798, x_361); +x_800 = l_Lean_Name_mkStr2(x_363, x_796); lean_inc(x_800); x_801 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_801, 0, x_800); -lean_ctor_set(x_801, 1, x_363); +lean_ctor_set(x_801, 1, x_345); x_802 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_802, 0, x_800); x_803 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_803, 0, x_802); -lean_ctor_set(x_803, 1, x_371); +lean_ctor_set(x_803, 1, x_356); x_804 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_804, 0, x_801); lean_ctor_set(x_804, 1, x_803); @@ -10198,17 +10198,17 @@ lean_ctor_set(x_805, 0, x_705); lean_ctor_set(x_805, 1, x_797); lean_ctor_set(x_805, 2, x_799); lean_ctor_set(x_805, 3, x_804); -lean_inc(x_351); +lean_inc(x_365); lean_inc(x_705); -x_806 = l_Lean_Syntax_node2(x_705, x_351, x_795, x_805); +x_806 = l_Lean_Syntax_node2(x_705, x_365, x_795, x_805); lean_inc(x_705); x_807 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_807, 0, x_705); -lean_ctor_set(x_807, 1, x_370); +lean_ctor_set(x_807, 1, x_362); lean_inc(x_705); -x_808 = l_Lean_Syntax_node6(x_705, x_367, x_760, x_761, x_793, x_794, x_806, x_807); +x_808 = l_Lean_Syntax_node6(x_705, x_339, x_760, x_761, x_793, x_794, x_806, x_807); lean_inc(x_705); -x_809 = l_Lean_Syntax_node1(x_705, x_351, x_808); +x_809 = l_Lean_Syntax_node1(x_705, x_365, x_808); x_810 = l_Lean_Syntax_node4(x_705, x_650, x_701, x_755, x_759, x_809); x_811 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_811, 0, x_754); @@ -10221,7 +10221,7 @@ goto block_16; block_855: { lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; -x_847 = 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__2_spec__2___lam__1(x_17, x_838, x_9, x_10); +x_847 = 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__2_spec__2___lam__1(x_17, x_843, x_9, x_10); x_848 = lean_ctor_get(x_847, 0); lean_inc(x_848); x_849 = lean_ctor_get(x_847, 1); @@ -10229,53 +10229,53 @@ lean_inc(x_849); lean_dec_ref(x_847); x_850 = l_Lean_mkIdentFrom(x_22, x_846, x_17); lean_dec(x_22); -lean_inc_ref(x_821); -lean_inc(x_826); +lean_inc_ref(x_832); +lean_inc(x_838); lean_inc(x_848); x_851 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_851, 0, x_848); -lean_ctor_set(x_851, 1, x_826); -lean_ctor_set(x_851, 2, x_821); +lean_ctor_set(x_851, 1, x_838); +lean_ctor_set(x_851, 2, x_832); if (lean_obj_tag(x_3) == 0) { lean_object* x_852; x_852 = 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__2_spec__2___closed__53; x_336 = x_813; -x_337 = x_817; -x_338 = x_816; -x_339 = x_815; -x_340 = x_814; +x_337 = x_814; +x_338 = x_815; +x_339 = x_816; +x_340 = x_817; x_341 = x_818; x_342 = x_819; -x_343 = x_820; -x_344 = x_848; -x_345 = x_821; -x_346 = x_822; -x_347 = x_823; +x_343 = x_821; +x_344 = x_820; +x_345 = x_822; +x_346 = x_823; +x_347 = x_849; x_348 = x_824; x_349 = x_825; -x_350 = x_850; -x_351 = x_826; -x_352 = x_827; -x_353 = x_828; -x_354 = x_829; -x_355 = x_830; -x_356 = x_831; +x_350 = x_826; +x_351 = x_827; +x_352 = x_828; +x_353 = x_829; +x_354 = x_830; +x_355 = x_831; +x_356 = x_833; x_357 = x_832; -x_358 = x_833; -x_359 = x_834; -x_360 = x_849; -x_361 = x_835; -x_362 = x_837; +x_358 = x_851; +x_359 = x_848; +x_360 = x_850; +x_361 = x_834; +x_362 = x_835; x_363 = x_836; -x_364 = x_838; -x_365 = x_851; +x_364 = x_837; +x_365 = x_838; x_366 = x_839; -x_367 = x_842; -x_368 = x_841; -x_369 = x_840; -x_370 = x_844; -x_371 = x_843; +x_367 = x_840; +x_368 = x_842; +x_369 = x_841; +x_370 = x_843; +x_371 = x_844; x_372 = x_845; x_373 = x_852; goto block_812; @@ -10287,41 +10287,41 @@ x_853 = lean_ctor_get(x_3, 0); lean_inc(x_853); x_854 = l_Array_mkArray1___redArg(x_853); x_336 = x_813; -x_337 = x_817; -x_338 = x_816; -x_339 = x_815; -x_340 = x_814; +x_337 = x_814; +x_338 = x_815; +x_339 = x_816; +x_340 = x_817; x_341 = x_818; x_342 = x_819; -x_343 = x_820; -x_344 = x_848; -x_345 = x_821; -x_346 = x_822; -x_347 = x_823; +x_343 = x_821; +x_344 = x_820; +x_345 = x_822; +x_346 = x_823; +x_347 = x_849; x_348 = x_824; x_349 = x_825; -x_350 = x_850; -x_351 = x_826; -x_352 = x_827; -x_353 = x_828; -x_354 = x_829; -x_355 = x_830; -x_356 = x_831; +x_350 = x_826; +x_351 = x_827; +x_352 = x_828; +x_353 = x_829; +x_354 = x_830; +x_355 = x_831; +x_356 = x_833; x_357 = x_832; -x_358 = x_833; -x_359 = x_834; -x_360 = x_849; -x_361 = x_835; -x_362 = x_837; +x_358 = x_851; +x_359 = x_848; +x_360 = x_850; +x_361 = x_834; +x_362 = x_835; x_363 = x_836; -x_364 = x_838; -x_365 = x_851; +x_364 = x_837; +x_365 = x_838; x_366 = x_839; -x_367 = x_842; -x_368 = x_841; -x_369 = x_840; -x_370 = x_844; -x_371 = x_843; +x_367 = x_840; +x_368 = x_842; +x_369 = x_841; +x_370 = x_843; +x_371 = x_844; x_372 = x_845; x_373 = x_854; goto block_812; @@ -10330,428 +10330,428 @@ goto block_812; block_1049: { lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; -lean_inc_ref(x_869); -x_874 = l_Array_append___redArg(x_869, x_873); +lean_inc_ref(x_862); +x_874 = l_Array_append___redArg(x_862, x_873); lean_dec_ref(x_873); -lean_inc(x_858); -lean_inc(x_871); +lean_inc(x_866); +lean_inc(x_865); x_875 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_875, 0, x_871); -lean_ctor_set(x_875, 1, x_858); +lean_ctor_set(x_875, 0, x_865); +lean_ctor_set(x_875, 1, x_866); lean_ctor_set(x_875, 2, x_874); -lean_inc_n(x_866, 6); -lean_inc(x_870); -lean_inc(x_871); -x_876 = l_Lean_Syntax_node7(x_871, x_870, x_866, x_866, x_875, x_866, x_866, x_866, x_866); +lean_inc_n(x_857, 6); +lean_inc(x_860); +lean_inc(x_865); +x_876 = l_Lean_Syntax_node7(x_865, x_860, x_857, x_857, x_875, x_857, x_857, x_857, x_857); x_877 = 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__4_spec__4___closed__25; -lean_inc_ref(x_859); -lean_inc_ref(x_861); +lean_inc_ref(x_868); lean_inc_ref(x_867); -x_878 = l_Lean_Name_mkStr4(x_867, x_861, x_859, x_877); +lean_inc_ref(x_870); +x_878 = l_Lean_Name_mkStr4(x_870, x_867, x_868, x_877); x_879 = 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__4_spec__4___closed__26; -lean_inc(x_871); +lean_inc(x_865); x_880 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_880, 0, x_871); +lean_ctor_set(x_880, 0, x_865); lean_ctor_set(x_880, 1, x_879); x_881 = l_Lake_configDecl___closed__8; -lean_inc_ref(x_859); -lean_inc_ref(x_861); +lean_inc_ref(x_868); lean_inc_ref(x_867); -x_882 = l_Lean_Name_mkStr4(x_867, x_861, x_859, x_881); -lean_inc(x_866); -lean_inc(x_862); +lean_inc_ref(x_870); +x_882 = l_Lean_Name_mkStr4(x_870, x_867, x_868, x_881); +lean_inc(x_857); +lean_inc(x_869); lean_inc(x_882); -lean_inc(x_871); -x_883 = l_Lean_Syntax_node2(x_871, x_882, x_862, x_866); +lean_inc(x_865); +x_883 = l_Lean_Syntax_node2(x_865, x_882, x_869, x_857); x_884 = 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__4_spec__4___closed__27; -lean_inc_ref(x_859); -lean_inc_ref(x_861); +lean_inc_ref(x_868); lean_inc_ref(x_867); -x_885 = l_Lean_Name_mkStr4(x_867, x_861, x_859, x_884); +lean_inc_ref(x_870); +x_885 = l_Lean_Name_mkStr4(x_870, x_867, x_868, x_884); x_886 = l_Lake_configDecl___closed__26; x_887 = 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__2_spec__2___closed__7; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_888 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_887); +lean_inc_ref(x_870); +x_888 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_887); x_889 = 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__2_spec__2___closed__8; -lean_inc(x_871); +lean_inc(x_865); x_890 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_890, 0, x_871); +lean_ctor_set(x_890, 0, x_865); lean_ctor_set(x_890, 1, x_889); x_891 = 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__2_spec__2___closed__9; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_892 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_891); +lean_inc_ref(x_870); +x_892 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_891); x_893 = 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__4_spec__4___closed__29; x_894 = 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__4_spec__4___closed__30; lean_inc(x_864); -lean_inc(x_863); -x_895 = l_Lean_addMacroScope(x_863, x_894, x_864); +lean_inc(x_861); +x_895 = l_Lean_addMacroScope(x_861, x_894, x_864); x_896 = l_Lake_configField___closed__1; x_897 = lean_box(0); x_898 = 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__4_spec__4___closed__35; -lean_inc(x_871); +lean_inc(x_865); x_899 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_899, 0, x_871); +lean_ctor_set(x_899, 0, x_865); lean_ctor_set(x_899, 1, x_893); lean_ctor_set(x_899, 2, x_895); lean_ctor_set(x_899, 3, x_898); lean_inc(x_24); lean_inc(x_1); -lean_inc(x_858); -lean_inc(x_871); -x_900 = l_Lean_Syntax_node2(x_871, x_858, x_1, x_24); +lean_inc(x_866); +lean_inc(x_865); +x_900 = l_Lean_Syntax_node2(x_865, x_866, x_1, x_24); lean_inc(x_892); -lean_inc(x_871); -x_901 = l_Lean_Syntax_node2(x_871, x_892, x_899, x_900); +lean_inc(x_865); +x_901 = l_Lean_Syntax_node2(x_865, x_892, x_899, x_900); lean_inc(x_888); -lean_inc(x_871); -x_902 = l_Lean_Syntax_node2(x_871, x_888, x_890, x_901); -lean_inc(x_858); -lean_inc(x_871); -x_903 = l_Lean_Syntax_node1(x_871, x_858, x_902); +lean_inc(x_865); +x_902 = l_Lean_Syntax_node2(x_865, x_888, x_890, x_901); lean_inc(x_866); -lean_inc(x_871); -x_904 = l_Lean_Syntax_node2(x_871, x_885, x_866, x_903); +lean_inc(x_865); +x_903 = l_Lean_Syntax_node1(x_865, x_866, x_902); +lean_inc(x_857); +lean_inc(x_865); +x_904 = l_Lean_Syntax_node2(x_865, x_885, x_857, x_903); x_905 = 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__4_spec__4___closed__36; -lean_inc_ref(x_859); -lean_inc_ref(x_861); +lean_inc_ref(x_868); lean_inc_ref(x_867); -x_906 = l_Lean_Name_mkStr4(x_867, x_861, x_859, x_905); +lean_inc_ref(x_870); +x_906 = l_Lean_Name_mkStr4(x_870, x_867, x_868, x_905); x_907 = 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__4_spec__4___closed__37; -lean_inc(x_871); +lean_inc(x_865); x_908 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_908, 0, x_871); +lean_ctor_set(x_908, 0, x_865); lean_ctor_set(x_908, 1, x_907); x_909 = 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__2_spec__2___closed__32; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_910 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_909); +lean_inc_ref(x_870); +x_910 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_909); x_911 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls_spec__1___closed__0; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_912 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_911); +lean_inc_ref(x_870); +x_912 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_911); x_913 = l___private_Lake_Config_Meta_0__Lake_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil__lake___lam__0___closed__1; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_914 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_913); +lean_inc_ref(x_870); +x_914 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_913); x_915 = 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__4_spec__4___closed__39; x_916 = 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__4_spec__4___closed__40; lean_inc(x_864); -lean_inc(x_863); -x_917 = l_Lean_addMacroScope(x_863, x_916, x_864); +lean_inc(x_861); +x_917 = l_Lean_addMacroScope(x_861, x_916, x_864); x_918 = 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__4_spec__4___closed__44; -lean_inc(x_871); +lean_inc(x_865); x_919 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_919, 0, x_871); +lean_ctor_set(x_919, 0, x_865); lean_ctor_set(x_919, 1, x_915); lean_ctor_set(x_919, 2, x_917); lean_ctor_set(x_919, 3, x_918); -lean_inc(x_866); +lean_inc(x_857); lean_inc(x_914); -lean_inc(x_871); -x_920 = l_Lean_Syntax_node2(x_871, x_914, x_919, x_866); +lean_inc(x_865); +x_920 = l_Lean_Syntax_node2(x_865, x_914, x_919, x_857); x_921 = 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__4_spec__4___closed__46; x_922 = 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__4_spec__4___closed__47; lean_inc(x_864); -lean_inc(x_863); -x_923 = l_Lean_addMacroScope(x_863, x_922, x_864); -lean_inc(x_871); +lean_inc(x_861); +x_923 = l_Lean_addMacroScope(x_861, x_922, x_864); +lean_inc(x_865); x_924 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_924, 0, x_871); +lean_ctor_set(x_924, 0, x_865); lean_ctor_set(x_924, 1, x_921); lean_ctor_set(x_924, 2, x_923); lean_ctor_set(x_924, 3, x_897); lean_inc_ref(x_924); -lean_inc(x_858); -lean_inc(x_871); -x_925 = l_Lean_Syntax_node1(x_871, x_858, x_924); +lean_inc(x_866); +lean_inc(x_865); +x_925 = l_Lean_Syntax_node1(x_865, x_866, x_924); x_926 = 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__2_spec__2___closed__36; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_927 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_926); +lean_inc_ref(x_870); +x_927 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_926); x_928 = 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__2_spec__2___closed__19; -lean_inc(x_871); +lean_inc(x_865); x_929 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_929, 0, x_871); +lean_ctor_set(x_929, 0, x_865); lean_ctor_set(x_929, 1, x_928); x_930 = 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__4_spec__4___closed__48; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_931 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_930); +lean_inc_ref(x_870); +x_931 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_930); x_932 = 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__4_spec__4___closed__49; -lean_inc(x_871); +lean_inc(x_865); x_933 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_933, 0, x_871); +lean_ctor_set(x_933, 0, x_865); lean_ctor_set(x_933, 1, x_932); lean_inc(x_22); lean_inc_ref(x_924); -lean_inc(x_871); -x_934 = l_Lean_Syntax_node3(x_871, x_931, x_924, x_933, x_22); +lean_inc(x_865); +x_934 = l_Lean_Syntax_node3(x_865, x_931, x_924, x_933, x_22); lean_inc(x_934); -lean_inc(x_866); +lean_inc(x_857); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_871); -x_935 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_934); -lean_inc(x_866); +lean_inc(x_865); +x_935 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_934); +lean_inc(x_857); lean_inc(x_925); -lean_inc(x_858); -lean_inc(x_871); -x_936 = l_Lean_Syntax_node3(x_871, x_858, x_925, x_866, x_935); +lean_inc(x_866); +lean_inc(x_865); +x_936 = l_Lean_Syntax_node3(x_865, x_866, x_925, x_857, x_935); lean_inc(x_912); -lean_inc(x_871); -x_937 = l_Lean_Syntax_node2(x_871, x_912, x_920, x_936); +lean_inc(x_865); +x_937 = l_Lean_Syntax_node2(x_865, x_912, x_920, x_936); x_938 = 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__4_spec__4___closed__51; x_939 = 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__4_spec__4___closed__52; lean_inc(x_864); -lean_inc(x_863); -x_940 = l_Lean_addMacroScope(x_863, x_939, x_864); +lean_inc(x_861); +x_940 = l_Lean_addMacroScope(x_861, x_939, x_864); x_941 = 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__4_spec__4___closed__56; -lean_inc(x_871); +lean_inc(x_865); x_942 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_942, 0, x_871); +lean_ctor_set(x_942, 0, x_865); lean_ctor_set(x_942, 1, x_938); lean_ctor_set(x_942, 2, x_940); lean_ctor_set(x_942, 3, x_941); -lean_inc(x_866); +lean_inc(x_857); lean_inc(x_914); -lean_inc(x_871); -x_943 = l_Lean_Syntax_node2(x_871, x_914, x_942, x_866); +lean_inc(x_865); +x_943 = l_Lean_Syntax_node2(x_865, x_914, x_942, x_857); x_944 = 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__4_spec__4___closed__58; x_945 = 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__4_spec__4___closed__59; lean_inc(x_864); -lean_inc(x_863); -x_946 = l_Lean_addMacroScope(x_863, x_945, x_864); -lean_inc(x_871); +lean_inc(x_861); +x_946 = l_Lean_addMacroScope(x_861, x_945, x_864); +lean_inc(x_865); x_947 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_947, 0, x_871); +lean_ctor_set(x_947, 0, x_865); lean_ctor_set(x_947, 1, x_944); lean_ctor_set(x_947, 2, x_946); lean_ctor_set(x_947, 3, x_897); lean_inc_ref(x_924); lean_inc_ref(x_947); -lean_inc(x_858); -lean_inc(x_871); -x_948 = l_Lean_Syntax_node2(x_871, x_858, x_947, x_924); +lean_inc(x_866); +lean_inc(x_865); +x_948 = l_Lean_Syntax_node2(x_865, x_866, x_947, x_924); x_949 = 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__2_spec__2___closed__30; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_950 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_949); +lean_inc_ref(x_870); +x_950 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_949); x_951 = 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__2_spec__2___closed__31; -lean_inc(x_871); +lean_inc(x_865); x_952 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_952, 0, x_871); +lean_ctor_set(x_952, 0, x_865); lean_ctor_set(x_952, 1, x_951); x_953 = 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__4_spec__4___closed__60; -lean_inc(x_871); +lean_inc(x_865); x_954 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_954, 0, x_871); +lean_ctor_set(x_954, 0, x_865); lean_ctor_set(x_954, 1, x_953); -lean_inc(x_858); -lean_inc(x_871); -x_955 = l_Lean_Syntax_node2(x_871, x_858, x_925, x_954); +lean_inc(x_866); +lean_inc(x_865); +x_955 = l_Lean_Syntax_node2(x_865, x_866, x_925, x_954); x_956 = lean_box(0); x_957 = l_Lean_SourceInfo_fromRef(x_956, x_17); -lean_inc_ref(x_869); -lean_inc(x_858); +lean_inc_ref(x_862); +lean_inc(x_866); lean_inc(x_957); x_958 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_958, 0, x_957); -lean_ctor_set(x_958, 1, x_858); -lean_ctor_set(x_958, 2, x_869); +lean_ctor_set(x_958, 1, x_866); +lean_ctor_set(x_958, 2, x_862); lean_inc(x_22); lean_inc(x_914); x_959 = l_Lean_Syntax_node2(x_957, x_914, x_22, x_958); -lean_inc(x_866); +lean_inc(x_857); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_871); -x_960 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_947); -lean_inc_n(x_866, 2); -lean_inc(x_858); -lean_inc(x_871); -x_961 = l_Lean_Syntax_node3(x_871, x_858, x_866, x_866, x_960); +lean_inc(x_865); +x_960 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_947); +lean_inc_n(x_857, 2); +lean_inc(x_866); +lean_inc(x_865); +x_961 = l_Lean_Syntax_node3(x_865, x_866, x_857, x_857, x_960); lean_inc(x_959); lean_inc(x_912); -lean_inc(x_871); -x_962 = l_Lean_Syntax_node2(x_871, x_912, x_959, x_961); -lean_inc(x_858); -lean_inc(x_871); -x_963 = l_Lean_Syntax_node1(x_871, x_858, x_962); +lean_inc(x_865); +x_962 = l_Lean_Syntax_node2(x_865, x_912, x_959, x_961); +lean_inc(x_866); +lean_inc(x_865); +x_963 = l_Lean_Syntax_node1(x_865, x_866, x_962); lean_inc(x_910); -lean_inc(x_871); -x_964 = l_Lean_Syntax_node1(x_871, x_910, x_963); +lean_inc(x_865); +x_964 = l_Lean_Syntax_node1(x_865, x_910, x_963); x_965 = 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__2_spec__2___closed__0; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_966 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_965); -lean_inc(x_866); +lean_inc_ref(x_870); +x_966 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_965); +lean_inc(x_857); lean_inc(x_966); -lean_inc(x_871); -x_967 = l_Lean_Syntax_node1(x_871, x_966, x_866); +lean_inc(x_865); +x_967 = l_Lean_Syntax_node1(x_865, x_966, x_857); x_968 = 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__2_spec__2___closed__4; -lean_inc(x_871); +lean_inc(x_865); x_969 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_969, 0, x_871); +lean_ctor_set(x_969, 0, x_865); lean_ctor_set(x_969, 1, x_968); lean_inc_ref(x_969); -lean_inc(x_866); +lean_inc(x_857); lean_inc(x_967); lean_inc(x_955); lean_inc_ref(x_952); lean_inc(x_950); -lean_inc(x_871); -x_970 = l_Lean_Syntax_node6(x_871, x_950, x_952, x_955, x_964, x_967, x_866, x_969); -lean_inc(x_866); +lean_inc(x_865); +x_970 = l_Lean_Syntax_node6(x_865, x_950, x_952, x_955, x_964, x_967, x_857, x_969); +lean_inc(x_857); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_871); -x_971 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_970); +lean_inc(x_865); +x_971 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_970); +lean_inc(x_857); lean_inc(x_866); -lean_inc(x_858); -lean_inc(x_871); -x_972 = l_Lean_Syntax_node3(x_871, x_858, x_948, x_866, x_971); +lean_inc(x_865); +x_972 = l_Lean_Syntax_node3(x_865, x_866, x_948, x_857, x_971); lean_inc(x_912); -lean_inc(x_871); -x_973 = l_Lean_Syntax_node2(x_871, x_912, x_943, x_972); +lean_inc(x_865); +x_973 = l_Lean_Syntax_node2(x_865, x_912, x_943, x_972); x_974 = 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__4_spec__4___closed__62; x_975 = 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__4_spec__4___closed__63; lean_inc(x_864); -lean_inc(x_863); -x_976 = l_Lean_addMacroScope(x_863, x_975, x_864); +lean_inc(x_861); +x_976 = l_Lean_addMacroScope(x_861, x_975, x_864); x_977 = 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__4_spec__4___closed__65; -lean_inc(x_871); +lean_inc(x_865); x_978 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_978, 0, x_871); +lean_ctor_set(x_978, 0, x_865); lean_ctor_set(x_978, 1, x_974); lean_ctor_set(x_978, 2, x_976); lean_ctor_set(x_978, 3, x_977); -lean_inc(x_866); +lean_inc(x_857); lean_inc(x_914); -lean_inc(x_871); -x_979 = l_Lean_Syntax_node2(x_871, x_914, x_978, x_866); +lean_inc(x_865); +x_979 = l_Lean_Syntax_node2(x_865, x_914, x_978, x_857); x_980 = 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__4_spec__4___closed__67; x_981 = 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__4_spec__4___closed__68; lean_inc(x_864); -lean_inc(x_863); -x_982 = l_Lean_addMacroScope(x_863, x_981, x_864); -lean_inc(x_871); +lean_inc(x_861); +x_982 = l_Lean_addMacroScope(x_861, x_981, x_864); +lean_inc(x_865); x_983 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_983, 0, x_871); +lean_ctor_set(x_983, 0, x_865); lean_ctor_set(x_983, 1, x_980); lean_ctor_set(x_983, 2, x_982); lean_ctor_set(x_983, 3, x_897); lean_inc_ref(x_983); -lean_inc(x_858); -lean_inc(x_871); -x_984 = l_Lean_Syntax_node2(x_871, x_858, x_983, x_924); -lean_inc(x_858); -lean_inc(x_871); -x_985 = l_Lean_Syntax_node1(x_871, x_858, x_934); -lean_inc(x_892); -lean_inc(x_871); -x_986 = l_Lean_Syntax_node2(x_871, x_892, x_983, x_985); lean_inc(x_866); +lean_inc(x_865); +x_984 = l_Lean_Syntax_node2(x_865, x_866, x_983, x_924); +lean_inc(x_866); +lean_inc(x_865); +x_985 = l_Lean_Syntax_node1(x_865, x_866, x_934); +lean_inc(x_892); +lean_inc(x_865); +x_986 = l_Lean_Syntax_node2(x_865, x_892, x_983, x_985); +lean_inc(x_857); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_871); -x_987 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_986); -lean_inc_n(x_866, 2); -lean_inc(x_858); -lean_inc(x_871); -x_988 = l_Lean_Syntax_node3(x_871, x_858, x_866, x_866, x_987); +lean_inc(x_865); +x_987 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_986); +lean_inc_n(x_857, 2); +lean_inc(x_866); +lean_inc(x_865); +x_988 = l_Lean_Syntax_node3(x_865, x_866, x_857, x_857, x_987); lean_inc(x_912); -lean_inc(x_871); -x_989 = l_Lean_Syntax_node2(x_871, x_912, x_959, x_988); -lean_inc(x_858); -lean_inc(x_871); -x_990 = l_Lean_Syntax_node1(x_871, x_858, x_989); -lean_inc(x_910); -lean_inc(x_871); -x_991 = l_Lean_Syntax_node1(x_871, x_910, x_990); +lean_inc(x_865); +x_989 = l_Lean_Syntax_node2(x_865, x_912, x_959, x_988); lean_inc(x_866); +lean_inc(x_865); +x_990 = l_Lean_Syntax_node1(x_865, x_866, x_989); +lean_inc(x_910); +lean_inc(x_865); +x_991 = l_Lean_Syntax_node1(x_865, x_910, x_990); +lean_inc(x_857); lean_inc(x_950); -lean_inc(x_871); -x_992 = l_Lean_Syntax_node6(x_871, x_950, x_952, x_955, x_991, x_967, x_866, x_969); -lean_inc(x_866); +lean_inc(x_865); +x_992 = l_Lean_Syntax_node6(x_865, x_950, x_952, x_955, x_991, x_967, x_857, x_969); +lean_inc(x_857); lean_inc_ref(x_929); lean_inc(x_927); -lean_inc(x_871); -x_993 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_992); +lean_inc(x_865); +x_993 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_992); +lean_inc(x_857); lean_inc(x_866); -lean_inc(x_858); -lean_inc(x_871); -x_994 = l_Lean_Syntax_node3(x_871, x_858, x_984, x_866, x_993); +lean_inc(x_865); +x_994 = l_Lean_Syntax_node3(x_865, x_866, x_984, x_857, x_993); lean_inc(x_912); -lean_inc(x_871); -x_995 = l_Lean_Syntax_node2(x_871, x_912, x_979, x_994); +lean_inc(x_865); +x_995 = l_Lean_Syntax_node2(x_865, x_912, x_979, x_994); x_996 = 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__4_spec__4___closed__70; x_997 = 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__4_spec__4___closed__71; lean_inc(x_864); -lean_inc(x_863); -x_998 = l_Lean_addMacroScope(x_863, x_997, x_864); -lean_inc(x_871); +lean_inc(x_861); +x_998 = l_Lean_addMacroScope(x_861, x_997, x_864); +lean_inc(x_865); x_999 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_999, 0, x_871); +lean_ctor_set(x_999, 0, x_865); lean_ctor_set(x_999, 1, x_996); lean_ctor_set(x_999, 2, x_998); lean_ctor_set(x_999, 3, x_897); -lean_inc(x_866); +lean_inc(x_857); lean_inc(x_914); -lean_inc(x_871); -x_1000 = l_Lean_Syntax_node2(x_871, x_914, x_999, x_866); +lean_inc(x_865); +x_1000 = l_Lean_Syntax_node2(x_865, x_914, x_999, x_857); x_1001 = 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__4_spec__4___closed__72; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_1002 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_1001); -lean_inc(x_871); +lean_inc_ref(x_870); +x_1002 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_1001); +lean_inc(x_865); x_1003 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1003, 0, x_871); +lean_ctor_set(x_1003, 0, x_865); lean_ctor_set(x_1003, 1, x_1001); x_1004 = 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__4_spec__4___closed__73; -lean_inc_ref(x_861); lean_inc_ref(x_867); -x_1005 = l_Lean_Name_mkStr4(x_867, x_861, x_886, x_1004); +lean_inc_ref(x_870); +x_1005 = l_Lean_Name_mkStr4(x_870, x_867, x_886, x_1004); lean_inc(x_2); -lean_inc(x_858); -lean_inc(x_871); -x_1006 = l_Lean_Syntax_node1(x_871, x_858, x_2); +lean_inc(x_866); +lean_inc(x_865); +x_1006 = l_Lean_Syntax_node1(x_865, x_866, x_2); x_1007 = 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__4_spec__4___closed__74; -lean_inc(x_871); +lean_inc(x_865); x_1008 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1008, 0, x_871); +lean_ctor_set(x_1008, 0, x_865); lean_ctor_set(x_1008, 1, x_1007); -lean_inc(x_866); -lean_inc(x_871); -x_1009 = l_Lean_Syntax_node4(x_871, x_1005, x_1006, x_866, x_1008, x_25); -lean_inc(x_871); -x_1010 = l_Lean_Syntax_node2(x_871, x_1002, x_1003, x_1009); -lean_inc(x_866); +lean_inc(x_857); +lean_inc(x_865); +x_1009 = l_Lean_Syntax_node4(x_865, x_1005, x_1006, x_857, x_1008, x_25); +lean_inc(x_865); +x_1010 = l_Lean_Syntax_node2(x_865, x_1002, x_1003, x_1009); +lean_inc(x_857); lean_inc(x_927); -lean_inc(x_871); -x_1011 = l_Lean_Syntax_node3(x_871, x_927, x_929, x_866, x_1010); -lean_inc_n(x_866, 2); -lean_inc(x_858); -lean_inc(x_871); -x_1012 = l_Lean_Syntax_node3(x_871, x_858, x_866, x_866, x_1011); +lean_inc(x_865); +x_1011 = l_Lean_Syntax_node3(x_865, x_927, x_929, x_857, x_1010); +lean_inc_n(x_857, 2); +lean_inc(x_866); +lean_inc(x_865); +x_1012 = l_Lean_Syntax_node3(x_865, x_866, x_857, x_857, x_1011); lean_inc(x_912); -lean_inc(x_871); -x_1013 = l_Lean_Syntax_node2(x_871, x_912, x_1000, x_1012); -lean_inc_n(x_866, 3); -lean_inc(x_858); -lean_inc(x_871); -x_1014 = l_Lean_Syntax_node7(x_871, x_858, x_937, x_866, x_973, x_866, x_995, x_866, x_1013); -lean_inc(x_910); -lean_inc(x_871); -x_1015 = l_Lean_Syntax_node1(x_871, x_910, x_1014); +lean_inc(x_865); +x_1013 = l_Lean_Syntax_node2(x_865, x_912, x_1000, x_1012); +lean_inc_n(x_857, 3); lean_inc(x_866); +lean_inc(x_865); +x_1014 = l_Lean_Syntax_node7(x_865, x_866, x_937, x_857, x_973, x_857, x_995, x_857, x_1013); +lean_inc(x_910); +lean_inc(x_865); +x_1015 = l_Lean_Syntax_node1(x_865, x_910, x_1014); +lean_inc(x_857); +lean_inc(x_865); +x_1016 = l_Lean_Syntax_node3(x_865, x_906, x_908, x_1015, x_857); +lean_inc(x_865); +x_1017 = l_Lean_Syntax_node5(x_865, x_878, x_880, x_883, x_904, x_1016, x_857); lean_inc(x_871); -x_1016 = l_Lean_Syntax_node3(x_871, x_906, x_908, x_1015, x_866); -lean_inc(x_871); -x_1017 = l_Lean_Syntax_node5(x_871, x_878, x_880, x_883, x_904, x_1016, x_866); -lean_inc(x_872); -x_1018 = l_Lean_Syntax_node2(x_871, x_872, x_876, x_1017); +x_1018 = l_Lean_Syntax_node2(x_865, x_871, x_876, x_1017); x_1019 = lean_array_push(x_18, x_1018); lean_inc(x_856); lean_inc(x_22); @@ -10764,39 +10764,39 @@ if (x_1021 == 0) { lean_object* x_1022; x_1022 = 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__2_spec__2___lam__0(x_4, x_856); -x_293 = x_927; -x_294 = x_857; -x_295 = x_1019; -x_296 = x_912; -x_297 = x_910; -x_298 = x_859; -x_299 = x_861; -x_300 = x_882; -x_301 = x_869; -x_302 = x_870; -x_303 = x_914; -x_304 = x_886; -x_305 = x_872; -x_306 = x_858; -x_307 = x_860; -x_308 = x_892; -x_309 = x_862; -x_310 = x_889; -x_311 = x_864; -x_312 = x_863; -x_313 = x_865; -x_314 = x_1020; -x_315 = x_966; -x_316 = x_897; -x_317 = x_867; -x_318 = x_868; -x_319 = x_928; -x_320 = x_888; -x_321 = x_951; -x_322 = x_950; -x_323 = x_897; -x_324 = x_968; -x_325 = x_896; +x_293 = x_858; +x_294 = x_859; +x_295 = x_860; +x_296 = x_950; +x_297 = x_863; +x_298 = x_912; +x_299 = x_867; +x_300 = x_1019; +x_301 = x_951; +x_302 = x_897; +x_303 = x_966; +x_304 = x_1020; +x_305 = x_892; +x_306 = x_910; +x_307 = x_914; +x_308 = x_882; +x_309 = x_889; +x_310 = x_861; +x_311 = x_886; +x_312 = x_862; +x_313 = x_897; +x_314 = x_864; +x_315 = x_968; +x_316 = x_896; +x_317 = x_868; +x_318 = x_866; +x_319 = x_869; +x_320 = x_870; +x_321 = x_928; +x_322 = x_871; +x_323 = x_872; +x_324 = x_888; +x_325 = x_927; x_326 = x_1022; goto block_335; } @@ -10812,39 +10812,39 @@ x_1025 = lean_ctor_get(x_1023, 0); x_1026 = 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__2_spec__2___lam__0(x_4, x_1025); lean_ctor_set(x_1023, 0, x_1026); x_1027 = l_Lean_MacroScopesView_review(x_1023); -x_293 = x_927; -x_294 = x_857; -x_295 = x_1019; -x_296 = x_912; -x_297 = x_910; -x_298 = x_859; -x_299 = x_861; -x_300 = x_882; -x_301 = x_869; -x_302 = x_870; -x_303 = x_914; -x_304 = x_886; -x_305 = x_872; -x_306 = x_858; -x_307 = x_860; -x_308 = x_892; -x_309 = x_862; -x_310 = x_889; -x_311 = x_864; -x_312 = x_863; -x_313 = x_865; -x_314 = x_1020; -x_315 = x_966; -x_316 = x_897; -x_317 = x_867; -x_318 = x_868; -x_319 = x_928; -x_320 = x_888; -x_321 = x_951; -x_322 = x_950; -x_323 = x_897; -x_324 = x_968; -x_325 = x_896; +x_293 = x_858; +x_294 = x_859; +x_295 = x_860; +x_296 = x_950; +x_297 = x_863; +x_298 = x_912; +x_299 = x_867; +x_300 = x_1019; +x_301 = x_951; +x_302 = x_897; +x_303 = x_966; +x_304 = x_1020; +x_305 = x_892; +x_306 = x_910; +x_307 = x_914; +x_308 = x_882; +x_309 = x_889; +x_310 = x_861; +x_311 = x_886; +x_312 = x_862; +x_313 = x_897; +x_314 = x_864; +x_315 = x_968; +x_316 = x_896; +x_317 = x_868; +x_318 = x_866; +x_319 = x_869; +x_320 = x_870; +x_321 = x_928; +x_322 = x_871; +x_323 = x_872; +x_324 = x_888; +x_325 = x_927; x_326 = x_1027; goto block_335; } @@ -10867,39 +10867,39 @@ lean_ctor_set(x_1033, 1, x_1029); lean_ctor_set(x_1033, 2, x_1030); lean_ctor_set(x_1033, 3, x_1031); x_1034 = l_Lean_MacroScopesView_review(x_1033); -x_293 = x_927; -x_294 = x_857; -x_295 = x_1019; -x_296 = x_912; -x_297 = x_910; -x_298 = x_859; -x_299 = x_861; -x_300 = x_882; -x_301 = x_869; -x_302 = x_870; -x_303 = x_914; -x_304 = x_886; -x_305 = x_872; -x_306 = x_858; -x_307 = x_860; -x_308 = x_892; -x_309 = x_862; -x_310 = x_889; -x_311 = x_864; -x_312 = x_863; -x_313 = x_865; -x_314 = x_1020; -x_315 = x_966; -x_316 = x_897; -x_317 = x_867; -x_318 = x_868; -x_319 = x_928; -x_320 = x_888; -x_321 = x_951; -x_322 = x_950; -x_323 = x_897; -x_324 = x_968; -x_325 = x_896; +x_293 = x_858; +x_294 = x_859; +x_295 = x_860; +x_296 = x_950; +x_297 = x_863; +x_298 = x_912; +x_299 = x_867; +x_300 = x_1019; +x_301 = x_951; +x_302 = x_897; +x_303 = x_966; +x_304 = x_1020; +x_305 = x_892; +x_306 = x_910; +x_307 = x_914; +x_308 = x_882; +x_309 = x_889; +x_310 = x_861; +x_311 = x_886; +x_312 = x_862; +x_313 = x_897; +x_314 = x_864; +x_315 = x_968; +x_316 = x_896; +x_317 = x_868; +x_318 = x_866; +x_319 = x_869; +x_320 = x_870; +x_321 = x_928; +x_322 = x_871; +x_323 = x_872; +x_324 = x_888; +x_325 = x_927; x_326 = x_1034; goto block_335; } @@ -10915,39 +10915,39 @@ if (x_1035 == 0) { lean_object* x_1036; x_1036 = 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__4_spec__4___lam__1(x_4, x_856); -x_813 = x_927; -x_814 = x_857; -x_815 = x_1019; -x_816 = x_912; -x_817 = x_910; -x_818 = x_859; -x_819 = x_861; -x_820 = x_882; -x_821 = x_869; -x_822 = x_870; -x_823 = x_914; -x_824 = x_886; -x_825 = x_872; -x_826 = x_858; -x_827 = x_860; -x_828 = x_892; -x_829 = x_862; -x_830 = x_889; -x_831 = x_864; -x_832 = x_863; -x_833 = x_865; -x_834 = x_1020; -x_835 = x_966; -x_836 = x_897; -x_837 = x_867; -x_838 = x_868; -x_839 = x_928; -x_840 = x_888; -x_841 = x_951; -x_842 = x_950; -x_843 = x_897; -x_844 = x_968; -x_845 = x_896; +x_813 = x_858; +x_814 = x_859; +x_815 = x_860; +x_816 = x_950; +x_817 = x_863; +x_818 = x_912; +x_819 = x_867; +x_820 = x_1019; +x_821 = x_951; +x_822 = x_897; +x_823 = x_966; +x_824 = x_1020; +x_825 = x_892; +x_826 = x_910; +x_827 = x_914; +x_828 = x_882; +x_829 = x_889; +x_830 = x_886; +x_831 = x_861; +x_832 = x_862; +x_833 = x_897; +x_834 = x_864; +x_835 = x_968; +x_836 = x_896; +x_837 = x_868; +x_838 = x_866; +x_839 = x_869; +x_840 = x_870; +x_841 = x_928; +x_842 = x_871; +x_843 = x_872; +x_844 = x_888; +x_845 = x_927; x_846 = x_1036; goto block_855; } @@ -10963,39 +10963,39 @@ x_1039 = lean_ctor_get(x_1037, 0); x_1040 = 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__4_spec__4___lam__1(x_4, x_1039); lean_ctor_set(x_1037, 0, x_1040); x_1041 = l_Lean_MacroScopesView_review(x_1037); -x_813 = x_927; -x_814 = x_857; -x_815 = x_1019; -x_816 = x_912; -x_817 = x_910; -x_818 = x_859; -x_819 = x_861; -x_820 = x_882; -x_821 = x_869; -x_822 = x_870; -x_823 = x_914; -x_824 = x_886; -x_825 = x_872; -x_826 = x_858; -x_827 = x_860; -x_828 = x_892; -x_829 = x_862; -x_830 = x_889; -x_831 = x_864; -x_832 = x_863; -x_833 = x_865; -x_834 = x_1020; -x_835 = x_966; -x_836 = x_897; -x_837 = x_867; -x_838 = x_868; -x_839 = x_928; -x_840 = x_888; -x_841 = x_951; -x_842 = x_950; -x_843 = x_897; -x_844 = x_968; -x_845 = x_896; +x_813 = x_858; +x_814 = x_859; +x_815 = x_860; +x_816 = x_950; +x_817 = x_863; +x_818 = x_912; +x_819 = x_867; +x_820 = x_1019; +x_821 = x_951; +x_822 = x_897; +x_823 = x_966; +x_824 = x_1020; +x_825 = x_892; +x_826 = x_910; +x_827 = x_914; +x_828 = x_882; +x_829 = x_889; +x_830 = x_886; +x_831 = x_861; +x_832 = x_862; +x_833 = x_897; +x_834 = x_864; +x_835 = x_968; +x_836 = x_896; +x_837 = x_868; +x_838 = x_866; +x_839 = x_869; +x_840 = x_870; +x_841 = x_928; +x_842 = x_871; +x_843 = x_872; +x_844 = x_888; +x_845 = x_927; x_846 = x_1041; goto block_855; } @@ -11018,39 +11018,39 @@ lean_ctor_set(x_1047, 1, x_1043); lean_ctor_set(x_1047, 2, x_1044); lean_ctor_set(x_1047, 3, x_1045); x_1048 = l_Lean_MacroScopesView_review(x_1047); -x_813 = x_927; -x_814 = x_857; -x_815 = x_1019; -x_816 = x_912; -x_817 = x_910; -x_818 = x_859; -x_819 = x_861; -x_820 = x_882; -x_821 = x_869; -x_822 = x_870; -x_823 = x_914; -x_824 = x_886; -x_825 = x_872; -x_826 = x_858; -x_827 = x_860; -x_828 = x_892; -x_829 = x_862; -x_830 = x_889; -x_831 = x_864; -x_832 = x_863; -x_833 = x_865; -x_834 = x_1020; -x_835 = x_966; -x_836 = x_897; -x_837 = x_867; -x_838 = x_868; -x_839 = x_928; -x_840 = x_888; -x_841 = x_951; -x_842 = x_950; -x_843 = x_897; -x_844 = x_968; -x_845 = x_896; +x_813 = x_858; +x_814 = x_859; +x_815 = x_860; +x_816 = x_950; +x_817 = x_863; +x_818 = x_912; +x_819 = x_867; +x_820 = x_1019; +x_821 = x_951; +x_822 = x_897; +x_823 = x_966; +x_824 = x_1020; +x_825 = x_892; +x_826 = x_910; +x_827 = x_914; +x_828 = x_882; +x_829 = x_889; +x_830 = x_886; +x_831 = x_861; +x_832 = x_862; +x_833 = x_897; +x_834 = x_864; +x_835 = x_968; +x_836 = x_896; +x_837 = x_868; +x_838 = x_866; +x_839 = x_869; +x_840 = x_870; +x_841 = x_928; +x_842 = x_871; +x_843 = x_872; +x_844 = x_888; +x_845 = x_927; x_846 = x_1048; goto block_855; } @@ -11085,27 +11085,27 @@ if (lean_obj_tag(x_3) == 0) lean_object* x_1067; x_1067 = 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__2_spec__2___closed__53; lean_inc(x_1056); -lean_inc(x_1051); lean_inc(x_1053); +lean_inc(x_1055); lean_inc(x_1052); +lean_inc(x_1051); lean_inc(x_1054); -lean_inc(x_1055); -x_857 = x_1055; -x_858 = x_1064; -x_859 = x_1061; -x_860 = x_1054; -x_861 = x_1060; -x_862 = x_1057; -x_863 = x_1052; +x_857 = x_1066; +x_858 = x_1054; +x_859 = x_1051; +x_860 = x_1063; +x_861 = x_1052; +x_862 = x_1065; +x_863 = x_1055; x_864 = x_1053; -x_865 = x_1051; -x_866 = x_1066; -x_867 = x_1059; -x_868 = x_1056; -x_869 = x_1065; -x_870 = x_1063; -x_871 = x_1058; -x_872 = x_1062; +x_865 = x_1058; +x_866 = x_1064; +x_867 = x_1060; +x_868 = x_1061; +x_869 = x_1057; +x_870 = x_1059; +x_871 = x_1062; +x_872 = x_1056; x_873 = x_1067; goto block_1049; } @@ -11116,27 +11116,27 @@ x_1068 = lean_ctor_get(x_3, 0); lean_inc(x_1068); x_1069 = l_Array_mkArray1___redArg(x_1068); lean_inc(x_1056); -lean_inc(x_1051); lean_inc(x_1053); +lean_inc(x_1055); lean_inc(x_1052); +lean_inc(x_1051); lean_inc(x_1054); -lean_inc(x_1055); -x_857 = x_1055; -x_858 = x_1064; -x_859 = x_1061; -x_860 = x_1054; -x_861 = x_1060; -x_862 = x_1057; -x_863 = x_1052; +x_857 = x_1066; +x_858 = x_1054; +x_859 = x_1051; +x_860 = x_1063; +x_861 = x_1052; +x_862 = x_1065; +x_863 = x_1055; x_864 = x_1053; -x_865 = x_1051; -x_866 = x_1066; -x_867 = x_1059; -x_868 = x_1056; -x_869 = x_1065; -x_870 = x_1063; -x_871 = x_1058; -x_872 = x_1062; +x_865 = x_1058; +x_866 = x_1064; +x_867 = x_1060; +x_868 = x_1061; +x_869 = x_1057; +x_870 = x_1059; +x_871 = x_1062; +x_872 = x_1056; x_873 = x_1069; goto block_1049; } @@ -11968,102 +11968,102 @@ goto block_563; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; x_55 = l_Array_append___redArg(x_25, x_54); lean_dec_ref(x_54); -lean_inc(x_30); +lean_inc(x_37); x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_30); +lean_ctor_set(x_56, 0, x_37); lean_ctor_set(x_56, 1, x_24); lean_ctor_set(x_56, 2, x_55); -lean_inc_n(x_48, 6); -lean_inc(x_30); -x_57 = l_Lean_Syntax_node7(x_30, x_31, x_48, x_48, x_56, x_48, x_48, x_48, x_48); -lean_inc(x_48); -lean_inc(x_30); -x_58 = l_Lean_Syntax_node1(x_30, x_53, x_48); -lean_inc(x_30); +lean_inc_n(x_47, 6); +lean_inc(x_37); +x_57 = l_Lean_Syntax_node7(x_37, x_35, x_47, x_47, x_56, x_47, x_47, x_47, x_47); +lean_inc(x_47); +lean_inc(x_37); +x_58 = l_Lean_Syntax_node1(x_37, x_39, x_47); +lean_inc(x_37); x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_30); -lean_ctor_set(x_59, 1, x_49); -lean_inc(x_48); -lean_inc(x_30); -x_60 = l_Lean_Syntax_node2(x_30, x_52, x_36, x_48); -lean_inc(x_30); -x_61 = l_Lean_Syntax_node1(x_30, x_24, x_60); -lean_inc(x_30); +lean_ctor_set(x_59, 0, x_37); +lean_ctor_set(x_59, 1, x_43); +lean_inc(x_47); +lean_inc(x_37); +x_60 = l_Lean_Syntax_node2(x_37, x_46, x_38, x_47); +lean_inc(x_37); +x_61 = l_Lean_Syntax_node1(x_37, x_24, x_60); +lean_inc(x_37); x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_30); -lean_ctor_set(x_62, 1, x_42); +lean_ctor_set(x_62, 0, x_37); +lean_ctor_set(x_62, 1, x_51); x_63 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__18; x_64 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__19; x_65 = l_Lean_addMacroScope(x_8, x_64, x_9); x_66 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__23; -lean_inc(x_30); +lean_inc(x_37); x_67 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_67, 0, x_30); +lean_ctor_set(x_67, 0, x_37); lean_ctor_set(x_67, 1, x_63); lean_ctor_set(x_67, 2, x_65); lean_ctor_set(x_67, 3, x_66); -lean_inc(x_30); -x_68 = l_Lean_Syntax_node1(x_30, x_24, x_4); -lean_inc(x_30); -x_69 = l_Lean_Syntax_node2(x_30, x_37, x_67, x_68); -lean_inc(x_30); -x_70 = l_Lean_Syntax_node2(x_30, x_34, x_62, x_69); -lean_inc(x_48); -lean_inc(x_30); -x_71 = l_Lean_Syntax_node2(x_30, x_50, x_48, x_70); -lean_inc(x_30); +lean_inc(x_37); +x_68 = l_Lean_Syntax_node1(x_37, x_24, x_4); +lean_inc(x_37); +x_69 = l_Lean_Syntax_node2(x_37, x_41, x_67, x_68); +lean_inc(x_37); +x_70 = l_Lean_Syntax_node2(x_37, x_53, x_62, x_69); +lean_inc(x_47); +lean_inc(x_37); +x_71 = l_Lean_Syntax_node2(x_37, x_32, x_47, x_70); +lean_inc(x_37); x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_30); -lean_ctor_set(x_72, 1, x_46); -lean_inc(x_30); +lean_ctor_set(x_72, 0, x_37); +lean_ctor_set(x_72, 1, x_31); +lean_inc(x_37); x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_30); -lean_ctor_set(x_73, 1, x_41); +lean_ctor_set(x_73, 0, x_37); +lean_ctor_set(x_73, 1, x_45); x_74 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__25; x_75 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__27; -lean_inc(x_30); +lean_inc(x_37); x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_30); +lean_ctor_set(x_76, 0, x_37); lean_ctor_set(x_76, 1, x_23); -lean_inc(x_30); +lean_inc(x_37); x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_30); +lean_ctor_set(x_77, 0, x_37); lean_ctor_set(x_77, 1, x_29); lean_inc_ref(x_77); lean_inc_ref(x_76); -lean_inc(x_30); -x_78 = l_Lean_Syntax_node2(x_30, x_75, x_76, x_77); -lean_inc(x_48); -lean_inc(x_30); -x_79 = l_Lean_Syntax_node1(x_30, x_26, x_48); -lean_inc(x_48); -lean_inc(x_30); -x_80 = l_Lean_Syntax_node1(x_30, x_28, x_48); -lean_inc_n(x_48, 2); -lean_inc(x_30); -x_81 = l_Lean_Syntax_node6(x_30, x_22, x_76, x_48, x_79, x_80, x_48, x_77); -lean_inc(x_30); -x_82 = l_Lean_Syntax_node2(x_30, x_74, x_78, x_81); -lean_inc(x_30); -x_83 = l_Lean_Syntax_node1(x_30, x_24, x_82); -lean_inc(x_30); +lean_inc(x_37); +x_78 = l_Lean_Syntax_node2(x_37, x_75, x_76, x_77); +lean_inc(x_47); +lean_inc(x_37); +x_79 = l_Lean_Syntax_node1(x_37, x_26, x_47); +lean_inc(x_47); +lean_inc(x_37); +x_80 = l_Lean_Syntax_node1(x_37, x_28, x_47); +lean_inc_n(x_47, 2); +lean_inc(x_37); +x_81 = l_Lean_Syntax_node6(x_37, x_22, x_76, x_47, x_79, x_80, x_47, x_77); +lean_inc(x_37); +x_82 = l_Lean_Syntax_node2(x_37, x_74, x_78, x_81); +lean_inc(x_37); +x_83 = l_Lean_Syntax_node1(x_37, x_24, x_82); +lean_inc(x_37); x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_30); -lean_ctor_set(x_84, 1, x_43); -lean_inc(x_30); -x_85 = l_Lean_Syntax_node3(x_30, x_40, x_73, x_83, x_84); -lean_inc_n(x_48, 2); -lean_inc(x_30); -x_86 = l_Lean_Syntax_node2(x_30, x_35, x_48, x_48); -lean_inc(x_48); -lean_inc(x_30); -x_87 = l_Lean_Syntax_node4(x_30, x_33, x_72, x_85, x_86, x_48); -lean_inc(x_30); -x_88 = l_Lean_Syntax_node6(x_30, x_47, x_58, x_59, x_48, x_61, x_71, x_87); -x_89 = l_Lean_Syntax_node2(x_30, x_32, x_57, x_88); -x_90 = lean_array_push(x_51, x_39); -x_91 = lean_array_push(x_90, x_45); -x_92 = lean_array_push(x_91, x_44); +lean_ctor_set(x_84, 0, x_37); +lean_ctor_set(x_84, 1, x_49); +lean_inc(x_37); +x_85 = l_Lean_Syntax_node3(x_37, x_36, x_73, x_83, x_84); +lean_inc_n(x_47, 2); +lean_inc(x_37); +x_86 = l_Lean_Syntax_node2(x_37, x_44, x_47, x_47); +lean_inc(x_47); +lean_inc(x_37); +x_87 = l_Lean_Syntax_node4(x_37, x_40, x_72, x_85, x_86, x_47); +lean_inc(x_37); +x_88 = l_Lean_Syntax_node6(x_37, x_48, x_58, x_59, x_47, x_61, x_71, x_87); +x_89 = l_Lean_Syntax_node2(x_37, x_50, x_57, x_88); +x_90 = lean_array_push(x_52, x_33); +x_91 = lean_array_push(x_90, x_34); +x_92 = lean_array_push(x_91, x_30); x_93 = lean_array_push(x_92, x_89); if (lean_is_scalar(x_14)) { x_94 = lean_alloc_ctor(0, 2, 0); @@ -12071,13 +12071,13 @@ if (lean_is_scalar(x_14)) { x_94 = x_14; } lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_38); +lean_ctor_set(x_94, 1, x_42); return x_94; } block_126: { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_118 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_97); +x_118 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_110); lean_dec_ref(x_6); lean_dec(x_10); x_119 = lean_ctor_get(x_118, 0); @@ -12096,30 +12096,30 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_123; x_123 = 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__2_spec__2___closed__53; -x_30 = x_119; -x_31 = x_96; +x_30 = x_96; +x_31 = x_97; x_32 = x_98; x_33 = x_99; x_34 = x_100; x_35 = x_101; -x_36 = x_121; -x_37 = x_102; -x_38 = x_120; +x_36 = x_102; +x_37 = x_119; +x_38 = x_121; x_39 = x_103; x_40 = x_104; x_41 = x_105; -x_42 = x_106; -x_43 = x_107; +x_42 = x_120; +x_43 = x_106; x_44 = x_108; -x_45 = x_109; -x_46 = x_110; -x_47 = x_111; -x_48 = x_122; -x_49 = x_113; -x_50 = x_112; -x_51 = x_115; -x_52 = x_114; -x_53 = x_116; +x_45 = x_107; +x_46 = x_109; +x_47 = x_122; +x_48 = x_112; +x_49 = x_111; +x_50 = x_114; +x_51 = x_113; +x_52 = x_116; +x_53 = x_115; x_54 = x_123; goto block_95; } @@ -12130,30 +12130,30 @@ x_124 = lean_ctor_get(x_1, 0); lean_inc(x_124); lean_dec_ref(x_1); x_125 = l_Array_mkArray1___redArg(x_124); -x_30 = x_119; -x_31 = x_96; +x_30 = x_96; +x_31 = x_97; x_32 = x_98; x_33 = x_99; x_34 = x_100; x_35 = x_101; -x_36 = x_121; -x_37 = x_102; -x_38 = x_120; +x_36 = x_102; +x_37 = x_119; +x_38 = x_121; x_39 = x_103; x_40 = x_104; x_41 = x_105; -x_42 = x_106; -x_43 = x_107; +x_42 = x_120; +x_43 = x_106; x_44 = x_108; -x_45 = x_109; -x_46 = x_110; -x_47 = x_111; -x_48 = x_122; -x_49 = x_113; -x_50 = x_112; -x_51 = x_115; -x_52 = x_114; -x_53 = x_116; +x_45 = x_107; +x_46 = x_109; +x_47 = x_122; +x_48 = x_112; +x_49 = x_111; +x_50 = x_114; +x_51 = x_113; +x_52 = x_116; +x_53 = x_115; x_54 = x_125; goto block_95; } @@ -12163,34 +12163,34 @@ goto block_95; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; x_155 = l_Array_append___redArg(x_25, x_154); lean_dec_ref(x_154); -lean_inc(x_143); +lean_inc(x_148); x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_143); +lean_ctor_set(x_156, 0, x_148); lean_ctor_set(x_156, 1, x_24); lean_ctor_set(x_156, 2, x_155); -lean_inc_n(x_128, 6); -lean_inc(x_127); -lean_inc(x_143); -x_157 = l_Lean_Syntax_node7(x_143, x_127, x_128, x_128, x_156, x_128, x_128, x_128, x_128); +lean_inc_n(x_145, 6); lean_inc(x_128); -lean_inc(x_153); -lean_inc(x_143); -x_158 = l_Lean_Syntax_node1(x_143, x_153, x_128); -lean_inc_ref(x_151); +lean_inc(x_148); +x_157 = l_Lean_Syntax_node7(x_148, x_128, x_145, x_145, x_156, x_145, x_145, x_145, x_145); +lean_inc(x_145); lean_inc(x_143); +lean_inc(x_148); +x_158 = l_Lean_Syntax_node1(x_148, x_143, x_145); +lean_inc_ref(x_144); +lean_inc(x_148); x_159 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_159, 0, x_143); -lean_ctor_set(x_159, 1, x_151); -lean_inc(x_128); -lean_inc(x_140); -lean_inc(x_143); -x_160 = l_Lean_Syntax_node2(x_143, x_140, x_149, x_128); -lean_inc(x_143); -x_161 = l_Lean_Syntax_node1(x_143, x_24, x_160); +lean_ctor_set(x_159, 0, x_148); +lean_ctor_set(x_159, 1, x_144); +lean_inc(x_145); +lean_inc(x_147); +lean_inc(x_148); +x_160 = l_Lean_Syntax_node2(x_148, x_147, x_151, x_145); +lean_inc(x_148); +x_161 = l_Lean_Syntax_node1(x_148, x_24, x_160); lean_inc_ref(x_136); -lean_inc(x_143); +lean_inc(x_148); x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_143); +lean_ctor_set(x_162, 0, x_148); lean_ctor_set(x_162, 1, x_136); x_163 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__28; x_164 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__29; @@ -12211,75 +12211,75 @@ lean_ctor_set(x_170, 1, x_17); x_171 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_171, 0, x_168); lean_ctor_set(x_171, 1, x_170); -lean_inc(x_143); +lean_inc(x_148); x_172 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_172, 0, x_143); +lean_ctor_set(x_172, 0, x_148); lean_ctor_set(x_172, 1, x_164); lean_ctor_set(x_172, 2, x_166); lean_ctor_set(x_172, 3, x_171); -lean_inc(x_143); -x_173 = l_Lean_Syntax_node1(x_143, x_24, x_147); +lean_inc(x_148); +x_173 = l_Lean_Syntax_node1(x_148, x_24, x_139); +lean_inc(x_129); +lean_inc(x_148); +x_174 = l_Lean_Syntax_node2(x_148, x_129, x_172, x_173); +lean_inc(x_153); +lean_inc(x_148); +x_175 = l_Lean_Syntax_node2(x_148, x_153, x_162, x_174); lean_inc(x_145); -lean_inc(x_143); -x_174 = l_Lean_Syntax_node2(x_143, x_145, x_172, x_173); -lean_inc(x_141); -lean_inc(x_143); -x_175 = l_Lean_Syntax_node2(x_143, x_141, x_162, x_174); -lean_inc(x_128); -lean_inc(x_139); -lean_inc(x_143); -x_176 = l_Lean_Syntax_node2(x_143, x_139, x_128, x_175); +lean_inc(x_127); +lean_inc(x_148); +x_176 = l_Lean_Syntax_node2(x_148, x_127, x_145, x_175); lean_inc_ref(x_138); -lean_inc(x_143); +lean_inc(x_148); x_177 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_177, 0, x_143); +lean_ctor_set(x_177, 0, x_148); lean_ctor_set(x_177, 1, x_138); -lean_inc_n(x_128, 2); -lean_inc(x_132); -lean_inc(x_143); -x_178 = l_Lean_Syntax_node2(x_143, x_132, x_128, x_128); -lean_inc(x_128); -lean_inc(x_142); -lean_inc(x_143); -x_179 = l_Lean_Syntax_node4(x_143, x_142, x_177, x_133, x_178, x_128); -lean_inc(x_150); -lean_inc(x_143); -x_180 = l_Lean_Syntax_node6(x_143, x_150, x_158, x_159, x_128, x_161, x_176, x_179); -lean_inc(x_129); -x_181 = l_Lean_Syntax_node2(x_143, x_129, x_157, x_180); -x_182 = l_Lean_Name_hasMacroScopes(x_144); +lean_inc_n(x_145, 2); +lean_inc(x_146); +lean_inc(x_148); +x_178 = l_Lean_Syntax_node2(x_148, x_146, x_145, x_145); +lean_inc(x_145); +lean_inc(x_130); +lean_inc(x_148); +x_179 = l_Lean_Syntax_node4(x_148, x_130, x_177, x_149, x_178, x_145); +lean_inc(x_135); +lean_inc(x_148); +x_180 = l_Lean_Syntax_node6(x_148, x_135, x_158, x_159, x_145, x_161, x_176, x_179); +lean_inc(x_137); +x_181 = l_Lean_Syntax_node2(x_148, x_137, x_157, x_180); +x_182 = l_Lean_Name_hasMacroScopes(x_132); if (x_182 == 0) { lean_object* x_183; -x_183 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__1(x_144); -x_96 = x_127; -x_97 = x_130; -x_98 = x_129; -x_99 = x_142; +x_183 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__1(x_132); +x_96 = x_181; +x_97 = x_138; +x_98 = x_127; +x_99 = x_140; x_100 = x_141; -x_101 = x_132; -x_102 = x_145; -x_103 = x_134; -x_104 = x_135; -x_105 = x_146; -x_106 = x_136; -x_107 = x_148; -x_108 = x_181; -x_109 = x_137; -x_110 = x_138; +x_101 = x_128; +x_102 = x_142; +x_103 = x_143; +x_104 = x_130; +x_105 = x_129; +x_106 = x_144; +x_107 = x_133; +x_108 = x_146; +x_109 = x_147; +x_110 = x_134; x_111 = x_150; -x_112 = x_139; -x_113 = x_151; -x_114 = x_140; -x_115 = x_152; -x_116 = x_153; +x_112 = x_135; +x_113 = x_136; +x_114 = x_137; +x_115 = x_153; +x_116 = x_152; x_117 = x_183; goto block_126; } else { lean_object* x_184; uint8_t x_185; -x_184 = l_Lean_extractMacroScopes(x_144); +x_184 = l_Lean_extractMacroScopes(x_132); x_185 = !lean_is_exclusive(x_184); if (x_185 == 0) { @@ -12288,27 +12288,27 @@ x_186 = lean_ctor_get(x_184, 0); x_187 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__1(x_186); lean_ctor_set(x_184, 0, x_187); x_188 = l_Lean_MacroScopesView_review(x_184); -x_96 = x_127; -x_97 = x_130; -x_98 = x_129; -x_99 = x_142; +x_96 = x_181; +x_97 = x_138; +x_98 = x_127; +x_99 = x_140; x_100 = x_141; -x_101 = x_132; -x_102 = x_145; -x_103 = x_134; -x_104 = x_135; -x_105 = x_146; -x_106 = x_136; -x_107 = x_148; -x_108 = x_181; -x_109 = x_137; -x_110 = x_138; +x_101 = x_128; +x_102 = x_142; +x_103 = x_143; +x_104 = x_130; +x_105 = x_129; +x_106 = x_144; +x_107 = x_133; +x_108 = x_146; +x_109 = x_147; +x_110 = x_134; x_111 = x_150; -x_112 = x_139; -x_113 = x_151; -x_114 = x_140; -x_115 = x_152; -x_116 = x_153; +x_112 = x_135; +x_113 = x_136; +x_114 = x_137; +x_115 = x_153; +x_116 = x_152; x_117 = x_188; goto block_126; } @@ -12331,27 +12331,27 @@ lean_ctor_set(x_194, 1, x_190); lean_ctor_set(x_194, 2, x_191); lean_ctor_set(x_194, 3, x_192); x_195 = l_Lean_MacroScopesView_review(x_194); -x_96 = x_127; -x_97 = x_130; -x_98 = x_129; -x_99 = x_142; +x_96 = x_181; +x_97 = x_138; +x_98 = x_127; +x_99 = x_140; x_100 = x_141; -x_101 = x_132; -x_102 = x_145; -x_103 = x_134; -x_104 = x_135; -x_105 = x_146; -x_106 = x_136; -x_107 = x_148; -x_108 = x_181; -x_109 = x_137; -x_110 = x_138; +x_101 = x_128; +x_102 = x_142; +x_103 = x_143; +x_104 = x_130; +x_105 = x_129; +x_106 = x_144; +x_107 = x_133; +x_108 = x_146; +x_109 = x_147; +x_110 = x_134; x_111 = x_150; -x_112 = x_139; -x_113 = x_151; -x_114 = x_140; -x_115 = x_152; -x_116 = x_153; +x_112 = x_135; +x_113 = x_136; +x_114 = x_137; +x_115 = x_153; +x_116 = x_152; x_117 = x_195; goto block_126; } @@ -12360,7 +12360,7 @@ goto block_126; block_354: { lean_object* x_221; uint8_t x_222; -x_221 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_210); +x_221 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_208); x_222 = !lean_is_exclusive(x_221); if (x_222 == 0) { @@ -12410,14 +12410,14 @@ lean_inc_ref(x_235); lean_inc(x_223); x_242 = l_Lean_Syntax_node2(x_223, x_237, x_241, x_235); x_243 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__41; -lean_inc_ref(x_213); +lean_inc_ref(x_197); lean_inc(x_223); lean_ctor_set_tag(x_221, 2); -lean_ctor_set(x_221, 1, x_213); +lean_ctor_set(x_221, 1, x_197); lean_inc_ref(x_235); lean_inc_ref(x_221); lean_inc(x_223); -x_244 = l_Lean_Syntax_node3(x_223, x_243, x_221, x_235, x_206); +x_244 = l_Lean_Syntax_node3(x_223, x_243, x_221, x_235, x_204); lean_inc_ref_n(x_235, 2); lean_inc(x_223); x_245 = l_Lean_Syntax_node3(x_223, x_24, x_235, x_235, x_244); @@ -12473,33 +12473,33 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_264; x_264 = 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__2_spec__2___closed__53; -x_127 = x_197; -x_128 = x_263; -x_129 = x_198; -x_130 = x_228; -x_131 = x_201; -x_132 = x_202; -x_133 = x_262; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_228; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_234; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_227; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_234; -x_148 = x_211; -x_149 = x_229; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_263; +x_146 = x_211; +x_147 = x_213; +x_148 = x_227; +x_149 = x_262; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_229; +x_152 = x_219; +x_153 = x_218; x_154 = x_264; goto block_196; } @@ -12509,33 +12509,33 @@ lean_object* x_265; lean_object* x_266; x_265 = lean_ctor_get(x_1, 0); lean_inc(x_265); x_266 = l_Array_mkArray1___redArg(x_265); -x_127 = x_197; -x_128 = x_263; -x_129 = x_198; -x_130 = x_228; -x_131 = x_201; -x_132 = x_202; -x_133 = x_262; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_228; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_234; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_227; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_234; -x_148 = x_211; -x_149 = x_229; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_263; +x_146 = x_211; +x_147 = x_213; +x_148 = x_227; +x_149 = x_262; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_229; +x_152 = x_219; +x_153 = x_218; x_154 = x_266; goto block_196; } @@ -12584,14 +12584,14 @@ lean_inc_ref(x_276); lean_inc(x_223); x_283 = l_Lean_Syntax_node2(x_223, x_278, x_282, x_276); x_284 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__41; -lean_inc_ref(x_213); +lean_inc_ref(x_197); lean_inc(x_223); lean_ctor_set_tag(x_221, 2); -lean_ctor_set(x_221, 1, x_213); +lean_ctor_set(x_221, 1, x_197); lean_inc_ref(x_276); lean_inc_ref(x_221); lean_inc(x_223); -x_285 = l_Lean_Syntax_node3(x_223, x_284, x_221, x_276, x_206); +x_285 = l_Lean_Syntax_node3(x_223, x_284, x_221, x_276, x_204); lean_inc_ref_n(x_276, 2); lean_inc(x_223); x_286 = l_Lean_Syntax_node3(x_223, x_24, x_276, x_276, x_285); @@ -12647,33 +12647,33 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_305; x_305 = 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__2_spec__2___closed__53; -x_127 = x_197; -x_128 = x_304; -x_129 = x_198; -x_130 = x_268; -x_131 = x_201; -x_132 = x_202; -x_133 = x_303; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_268; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_274; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_267; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_274; -x_148 = x_211; -x_149 = x_269; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_304; +x_146 = x_211; +x_147 = x_213; +x_148 = x_267; +x_149 = x_303; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_269; +x_152 = x_219; +x_153 = x_218; x_154 = x_305; goto block_196; } @@ -12683,33 +12683,33 @@ lean_object* x_306; lean_object* x_307; x_306 = lean_ctor_get(x_1, 0); lean_inc(x_306); x_307 = l_Array_mkArray1___redArg(x_306); -x_127 = x_197; -x_128 = x_304; -x_129 = x_198; -x_130 = x_268; -x_131 = x_201; -x_132 = x_202; -x_133 = x_303; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_268; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_274; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_267; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_274; -x_148 = x_211; -x_149 = x_269; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_304; +x_146 = x_211; +x_147 = x_213; +x_148 = x_267; +x_149 = x_303; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_269; +x_152 = x_219; +x_153 = x_218; x_154 = x_307; goto block_196; } @@ -12777,15 +12777,15 @@ lean_inc_ref(x_321); lean_inc(x_308); x_328 = l_Lean_Syntax_node2(x_308, x_323, x_327, x_321); x_329 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__41; -lean_inc_ref(x_213); +lean_inc_ref(x_197); lean_inc(x_308); x_330 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_330, 0, x_308); -lean_ctor_set(x_330, 1, x_213); +lean_ctor_set(x_330, 1, x_197); lean_inc_ref(x_321); lean_inc_ref(x_330); lean_inc(x_308); -x_331 = l_Lean_Syntax_node3(x_308, x_329, x_330, x_321, x_206); +x_331 = l_Lean_Syntax_node3(x_308, x_329, x_330, x_321, x_204); lean_inc_ref_n(x_321, 2); lean_inc(x_308); x_332 = l_Lean_Syntax_node3(x_308, x_24, x_321, x_321, x_331); @@ -12841,33 +12841,33 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_351; x_351 = 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__2_spec__2___closed__53; -x_127 = x_197; -x_128 = x_350; -x_129 = x_198; -x_130 = x_312; -x_131 = x_201; -x_132 = x_202; -x_133 = x_349; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_312; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_319; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_311; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_319; -x_148 = x_211; -x_149 = x_314; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_350; +x_146 = x_211; +x_147 = x_213; +x_148 = x_311; +x_149 = x_349; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_314; +x_152 = x_219; +x_153 = x_218; x_154 = x_351; goto block_196; } @@ -12877,33 +12877,33 @@ lean_object* x_352; lean_object* x_353; x_352 = lean_ctor_get(x_1, 0); lean_inc(x_352); x_353 = l_Array_mkArray1___redArg(x_352); -x_127 = x_197; -x_128 = x_350; -x_129 = x_198; -x_130 = x_312; -x_131 = x_201; -x_132 = x_202; -x_133 = x_349; -x_134 = x_205; -x_135 = x_207; -x_136 = x_209; -x_137 = x_212; -x_138 = x_213; -x_139 = x_216; -x_140 = x_218; +x_127 = x_198; +x_128 = x_201; +x_129 = x_205; +x_130 = x_206; +x_131 = x_207; +x_132 = x_210; +x_133 = x_212; +x_134 = x_312; +x_135 = x_215; +x_136 = x_216; +x_137 = x_217; +x_138 = x_197; +x_139 = x_319; +x_140 = x_199; x_141 = x_200; -x_142 = x_199; -x_143 = x_311; -x_144 = x_203; -x_145 = x_204; -x_146 = x_208; -x_147 = x_319; -x_148 = x_211; -x_149 = x_314; +x_142 = x_202; +x_143 = x_203; +x_144 = x_209; +x_145 = x_350; +x_146 = x_211; +x_147 = x_213; +x_148 = x_311; +x_149 = x_349; x_150 = x_214; -x_151 = x_215; -x_152 = x_217; -x_153 = x_219; +x_151 = x_314; +x_152 = x_219; +x_153 = x_218; x_154 = x_353; goto block_196; } @@ -12919,29 +12919,29 @@ x_372 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_372, 0, x_364); lean_ctor_set(x_372, 1, x_24); lean_ctor_set(x_372, 2, x_371); -lean_inc_n(x_366, 6); -lean_inc(x_355); +lean_inc_n(x_365, 6); +lean_inc(x_357); lean_inc(x_364); -x_373 = l_Lean_Syntax_node7(x_364, x_355, x_366, x_366, x_372, x_366, x_366, x_366, x_366); +x_373 = l_Lean_Syntax_node7(x_364, x_357, x_365, x_365, x_372, x_365, x_365, x_365, x_365); x_374 = 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__2_spec__2___closed__5; -lean_inc_ref(x_367); -x_375 = l_Lean_Name_mkStr4(x_19, x_20, x_367, x_374); +lean_inc_ref(x_358); +x_375 = l_Lean_Name_mkStr4(x_19, x_20, x_358, x_374); x_376 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__45; -lean_inc(x_366); +lean_inc(x_365); lean_inc(x_364); -x_377 = l_Lean_Syntax_node1(x_364, x_376, x_366); +x_377 = l_Lean_Syntax_node1(x_364, x_376, x_365); lean_inc(x_364); x_378 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_378, 0, x_364); lean_ctor_set(x_378, 1, x_374); -lean_inc(x_366); -lean_inc(x_369); +lean_inc(x_365); +lean_inc(x_367); lean_inc(x_364); -x_379 = l_Lean_Syntax_node2(x_364, x_369, x_362, x_366); +x_379 = l_Lean_Syntax_node2(x_364, x_367, x_361, x_365); lean_inc(x_364); x_380 = l_Lean_Syntax_node1(x_364, x_24, x_379); x_381 = l_Lake_configField___closed__19; -x_382 = l_Lean_Name_mkStr4(x_19, x_20, x_367, x_381); +x_382 = l_Lean_Name_mkStr4(x_19, x_20, x_358, x_381); x_383 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__46; x_384 = 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__2_spec__2___closed__8; lean_inc(x_364); @@ -12969,24 +12969,24 @@ lean_inc(x_364); x_394 = l_Lean_Syntax_node2(x_364, x_386, x_392, x_393); lean_inc(x_364); x_395 = l_Lean_Syntax_node2(x_364, x_383, x_385, x_394); -lean_inc(x_366); +lean_inc(x_365); lean_inc(x_382); lean_inc(x_364); -x_396 = l_Lean_Syntax_node2(x_364, x_382, x_366, x_395); -lean_inc_ref(x_365); +x_396 = l_Lean_Syntax_node2(x_364, x_382, x_365, x_395); +lean_inc_ref(x_355); lean_inc(x_364); x_397 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_397, 0, x_364); -lean_ctor_set(x_397, 1, x_365); +lean_ctor_set(x_397, 1, x_355); x_398 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__55; x_399 = 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__2_spec__2___closed__21; lean_inc(x_364); x_400 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_400, 0, x_364); lean_ctor_set(x_400, 1, x_399); -lean_inc(x_361); +lean_inc(x_359); lean_inc(x_364); -x_401 = l_Lean_Syntax_node1(x_364, x_24, x_361); +x_401 = l_Lean_Syntax_node1(x_364, x_24, x_359); x_402 = 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__2_spec__2___closed__22; lean_inc(x_364); x_403 = lean_alloc_ctor(2, 2, 0); @@ -12994,56 +12994,56 @@ lean_ctor_set(x_403, 0, x_364); lean_ctor_set(x_403, 1, x_402); lean_inc(x_364); x_404 = l_Lean_Syntax_node3(x_364, x_398, x_400, x_401, x_403); -lean_inc_n(x_366, 2); -lean_inc(x_358); -lean_inc(x_364); -x_405 = l_Lean_Syntax_node2(x_364, x_358, x_366, x_366); +lean_inc_n(x_365, 2); lean_inc(x_366); -lean_inc(x_357); lean_inc(x_364); -x_406 = l_Lean_Syntax_node4(x_364, x_357, x_397, x_404, x_405, x_366); +x_405 = l_Lean_Syntax_node2(x_364, x_366, x_365, x_365); +lean_inc(x_365); +lean_inc(x_360); +lean_inc(x_364); +x_406 = l_Lean_Syntax_node4(x_364, x_360, x_397, x_404, x_405, x_365); lean_inc(x_375); lean_inc(x_364); -x_407 = l_Lean_Syntax_node6(x_364, x_375, x_377, x_378, x_366, x_380, x_396, x_406); -lean_inc(x_356); -x_408 = l_Lean_Syntax_node2(x_364, x_356, x_373, x_407); -x_409 = l_Lean_Name_hasMacroScopes(x_359); +x_407 = l_Lean_Syntax_node6(x_364, x_375, x_377, x_378, x_365, x_380, x_396, x_406); +lean_inc(x_368); +x_408 = l_Lean_Syntax_node2(x_364, x_368, x_373, x_407); +x_409 = l_Lean_Name_hasMacroScopes(x_363); if (x_409 == 0) { lean_object* x_410; -lean_inc(x_359); -x_410 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__2(x_359); +lean_inc(x_363); +x_410 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__2(x_363); x_197 = x_355; -x_198 = x_356; -x_199 = x_357; -x_200 = x_383; -x_201 = x_390; -x_202 = x_358; -x_203 = x_359; -x_204 = x_386; -x_205 = x_360; -x_206 = x_361; -x_207 = x_398; -x_208 = x_399; -x_209 = x_384; +x_198 = x_382; +x_199 = x_356; +x_200 = x_408; +x_201 = x_357; +x_202 = x_398; +x_203 = x_376; +x_204 = x_359; +x_205 = x_386; +x_206 = x_360; +x_207 = x_390; +x_208 = x_362; +x_209 = x_374; x_210 = x_363; -x_211 = x_402; -x_212 = x_408; -x_213 = x_365; -x_214 = x_375; -x_215 = x_374; -x_216 = x_382; +x_211 = x_366; +x_212 = x_399; +x_213 = x_367; +x_214 = x_402; +x_215 = x_375; +x_216 = x_384; x_217 = x_368; -x_218 = x_369; -x_219 = x_376; +x_218 = x_383; +x_219 = x_369; x_220 = x_410; goto block_354; } else { lean_object* x_411; uint8_t x_412; -lean_inc(x_359); -x_411 = l_Lean_extractMacroScopes(x_359); +lean_inc(x_363); +x_411 = l_Lean_extractMacroScopes(x_363); x_412 = !lean_is_exclusive(x_411); if (x_412 == 0) { @@ -13053,28 +13053,28 @@ x_414 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__2(x_413); lean_ctor_set(x_411, 0, x_414); x_415 = l_Lean_MacroScopesView_review(x_411); x_197 = x_355; -x_198 = x_356; -x_199 = x_357; -x_200 = x_383; -x_201 = x_390; -x_202 = x_358; -x_203 = x_359; -x_204 = x_386; -x_205 = x_360; -x_206 = x_361; -x_207 = x_398; -x_208 = x_399; -x_209 = x_384; +x_198 = x_382; +x_199 = x_356; +x_200 = x_408; +x_201 = x_357; +x_202 = x_398; +x_203 = x_376; +x_204 = x_359; +x_205 = x_386; +x_206 = x_360; +x_207 = x_390; +x_208 = x_362; +x_209 = x_374; x_210 = x_363; -x_211 = x_402; -x_212 = x_408; -x_213 = x_365; -x_214 = x_375; -x_215 = x_374; -x_216 = x_382; +x_211 = x_366; +x_212 = x_399; +x_213 = x_367; +x_214 = x_402; +x_215 = x_375; +x_216 = x_384; x_217 = x_368; -x_218 = x_369; -x_219 = x_376; +x_218 = x_383; +x_219 = x_369; x_220 = x_415; goto block_354; } @@ -13098,28 +13098,28 @@ lean_ctor_set(x_421, 2, x_418); lean_ctor_set(x_421, 3, x_419); x_422 = l_Lean_MacroScopesView_review(x_421); x_197 = x_355; -x_198 = x_356; -x_199 = x_357; -x_200 = x_383; -x_201 = x_390; -x_202 = x_358; -x_203 = x_359; -x_204 = x_386; -x_205 = x_360; -x_206 = x_361; -x_207 = x_398; -x_208 = x_399; -x_209 = x_384; +x_198 = x_382; +x_199 = x_356; +x_200 = x_408; +x_201 = x_357; +x_202 = x_398; +x_203 = x_376; +x_204 = x_359; +x_205 = x_386; +x_206 = x_360; +x_207 = x_390; +x_208 = x_362; +x_209 = x_374; x_210 = x_363; -x_211 = x_402; -x_212 = x_408; -x_213 = x_365; -x_214 = x_375; -x_215 = x_374; -x_216 = x_382; +x_211 = x_366; +x_212 = x_399; +x_213 = x_367; +x_214 = x_402; +x_215 = x_375; +x_216 = x_384; x_217 = x_368; -x_218 = x_369; -x_219 = x_376; +x_218 = x_383; +x_219 = x_369; x_220 = x_422; goto block_354; } @@ -13128,7 +13128,7 @@ goto block_354; block_445: { lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_437 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_435); +x_437 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_432); x_438 = lean_ctor_get(x_437, 0); lean_inc(x_438); x_439 = lean_ctor_get(x_437, 1); @@ -13144,21 +13144,21 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_442; x_442 = 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__2_spec__2___closed__53; -x_355 = x_424; -x_356 = x_425; -x_357 = x_426; -x_358 = x_427; -x_359 = x_428; -x_360 = x_431; -x_361 = x_432; -x_362 = x_440; -x_363 = x_439; +x_355 = x_425; +x_356 = x_426; +x_357 = x_427; +x_358 = x_431; +x_359 = x_434; +x_360 = x_424; +x_361 = x_440; +x_362 = x_439; +x_363 = x_428; x_364 = x_438; -x_365 = x_429; -x_366 = x_441; +x_365 = x_441; +x_366 = x_429; x_367 = x_430; -x_368 = x_434; -x_369 = x_433; +x_368 = x_433; +x_369 = x_435; x_370 = x_442; goto block_423; } @@ -13168,21 +13168,21 @@ lean_object* x_443; lean_object* x_444; x_443 = lean_ctor_get(x_1, 0); lean_inc(x_443); x_444 = l_Array_mkArray1___redArg(x_443); -x_355 = x_424; -x_356 = x_425; -x_357 = x_426; -x_358 = x_427; -x_359 = x_428; -x_360 = x_431; -x_361 = x_432; -x_362 = x_440; -x_363 = x_439; +x_355 = x_425; +x_356 = x_426; +x_357 = x_427; +x_358 = x_431; +x_359 = x_434; +x_360 = x_424; +x_361 = x_440; +x_362 = x_439; +x_363 = x_428; x_364 = x_438; -x_365 = x_429; -x_366 = x_441; +x_365 = x_441; +x_366 = x_429; x_367 = x_430; -x_368 = x_434; -x_369 = x_433; +x_368 = x_433; +x_369 = x_435; x_370 = x_444; goto block_423; } @@ -13198,83 +13198,83 @@ x_457 = lean_ctor_get(x_449, 0); x_458 = lean_ctor_get(x_449, 1); x_459 = l_Array_append___redArg(x_25, x_455); lean_dec_ref(x_455); -lean_inc(x_454); +lean_inc(x_452); x_460 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_460, 0, x_454); +lean_ctor_set(x_460, 0, x_452); lean_ctor_set(x_460, 1, x_24); lean_ctor_set(x_460, 2, x_459); -lean_inc_n(x_450, 6); +lean_inc_n(x_448, 6); lean_inc(x_446); -lean_inc(x_454); -x_461 = l_Lean_Syntax_node7(x_454, x_446, x_450, x_450, x_460, x_450, x_450, x_450, x_450); +lean_inc(x_452); +x_461 = l_Lean_Syntax_node7(x_452, x_446, x_448, x_448, x_460, x_448, x_448, x_448, x_448); x_462 = 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__4_spec__4___closed__25; lean_inc_ref(x_451); x_463 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_462); x_464 = 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__4_spec__4___closed__26; -lean_inc(x_454); +lean_inc(x_452); lean_ctor_set_tag(x_449, 2); lean_ctor_set(x_449, 1, x_464); -lean_ctor_set(x_449, 0, x_454); +lean_ctor_set(x_449, 0, x_452); x_465 = l_Lake_configDecl___closed__8; lean_inc_ref(x_451); x_466 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_465); -lean_inc(x_450); -lean_inc(x_452); -lean_inc(x_466); +lean_inc(x_448); lean_inc(x_454); -x_467 = l_Lean_Syntax_node2(x_454, x_466, x_452, x_450); +lean_inc(x_466); +lean_inc(x_452); +x_467 = l_Lean_Syntax_node2(x_452, x_466, x_454, x_448); x_468 = 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__4_spec__4___closed__27; lean_inc_ref(x_451); x_469 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_468); -lean_inc_n(x_450, 2); -lean_inc(x_454); -x_470 = l_Lean_Syntax_node2(x_454, x_469, x_450, x_450); +lean_inc_n(x_448, 2); +lean_inc(x_452); +x_470 = l_Lean_Syntax_node2(x_452, x_469, x_448, x_448); x_471 = 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__2_spec__2___closed__18; lean_inc_ref(x_451); x_472 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_471); x_473 = 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__2_spec__2___closed__19; -lean_inc(x_454); +lean_inc(x_452); x_474 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_474, 0, x_454); +lean_ctor_set(x_474, 0, x_452); lean_ctor_set(x_474, 1, x_473); x_475 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__56; -lean_inc_n(x_450, 2); -lean_inc(x_454); -x_476 = l_Lean_Syntax_node2(x_454, x_475, x_450, x_450); -lean_inc(x_450); +lean_inc_n(x_448, 2); +lean_inc(x_452); +x_476 = l_Lean_Syntax_node2(x_452, x_475, x_448, x_448); +lean_inc(x_448); lean_inc(x_472); -lean_inc(x_454); -x_477 = l_Lean_Syntax_node4(x_454, x_472, x_474, x_458, x_476, x_450); -lean_inc(x_454); -x_478 = l_Lean_Syntax_node5(x_454, x_463, x_449, x_467, x_470, x_477, x_450); -lean_inc(x_447); -x_479 = l_Lean_Syntax_node2(x_454, x_447, x_461, x_478); -x_480 = l_Lean_Name_hasMacroScopes(x_448); +lean_inc(x_452); +x_477 = l_Lean_Syntax_node4(x_452, x_472, x_474, x_458, x_476, x_448); +lean_inc(x_452); +x_478 = l_Lean_Syntax_node5(x_452, x_463, x_449, x_467, x_470, x_477, x_448); +lean_inc(x_453); +x_479 = l_Lean_Syntax_node2(x_452, x_453, x_461, x_478); +x_480 = l_Lean_Name_hasMacroScopes(x_447); if (x_480 == 0) { lean_object* x_481; -lean_inc(x_448); -x_481 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__3(x_448); -x_424 = x_446; -x_425 = x_447; -x_426 = x_472; -x_427 = x_475; -x_428 = x_448; -x_429 = x_473; -x_430 = x_451; -x_431 = x_479; -x_432 = x_452; -x_433 = x_466; -x_434 = x_457; -x_435 = x_453; +lean_inc(x_447); +x_481 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__3(x_447); +x_424 = x_472; +x_425 = x_473; +x_426 = x_479; +x_427 = x_446; +x_428 = x_447; +x_429 = x_475; +x_430 = x_466; +x_431 = x_451; +x_432 = x_450; +x_433 = x_453; +x_434 = x_454; +x_435 = x_457; x_436 = x_481; goto block_445; } else { lean_object* x_482; uint8_t x_483; -lean_inc(x_448); -x_482 = l_Lean_extractMacroScopes(x_448); +lean_inc(x_447); +x_482 = l_Lean_extractMacroScopes(x_447); x_483 = !lean_is_exclusive(x_482); if (x_483 == 0) { @@ -13283,18 +13283,18 @@ x_484 = lean_ctor_get(x_482, 0); x_485 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__3(x_484); lean_ctor_set(x_482, 0, x_485); x_486 = l_Lean_MacroScopesView_review(x_482); -x_424 = x_446; -x_425 = x_447; -x_426 = x_472; -x_427 = x_475; -x_428 = x_448; -x_429 = x_473; -x_430 = x_451; -x_431 = x_479; -x_432 = x_452; -x_433 = x_466; -x_434 = x_457; -x_435 = x_453; +x_424 = x_472; +x_425 = x_473; +x_426 = x_479; +x_427 = x_446; +x_428 = x_447; +x_429 = x_475; +x_430 = x_466; +x_431 = x_451; +x_432 = x_450; +x_433 = x_453; +x_434 = x_454; +x_435 = x_457; x_436 = x_486; goto block_445; } @@ -13317,18 +13317,18 @@ lean_ctor_set(x_492, 1, x_488); lean_ctor_set(x_492, 2, x_489); lean_ctor_set(x_492, 3, x_490); x_493 = l_Lean_MacroScopesView_review(x_492); -x_424 = x_446; -x_425 = x_447; -x_426 = x_472; -x_427 = x_475; -x_428 = x_448; -x_429 = x_473; -x_430 = x_451; -x_431 = x_479; -x_432 = x_452; -x_433 = x_466; -x_434 = x_457; -x_435 = x_453; +x_424 = x_472; +x_425 = x_473; +x_426 = x_479; +x_427 = x_446; +x_428 = x_447; +x_429 = x_475; +x_430 = x_466; +x_431 = x_451; +x_432 = x_450; +x_433 = x_453; +x_434 = x_454; +x_435 = x_457; x_436 = x_493; goto block_445; } @@ -13344,83 +13344,83 @@ lean_inc(x_494); lean_dec(x_449); x_496 = l_Array_append___redArg(x_25, x_455); lean_dec_ref(x_455); -lean_inc(x_454); +lean_inc(x_452); x_497 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_497, 0, x_454); +lean_ctor_set(x_497, 0, x_452); lean_ctor_set(x_497, 1, x_24); lean_ctor_set(x_497, 2, x_496); -lean_inc_n(x_450, 6); +lean_inc_n(x_448, 6); lean_inc(x_446); -lean_inc(x_454); -x_498 = l_Lean_Syntax_node7(x_454, x_446, x_450, x_450, x_497, x_450, x_450, x_450, x_450); +lean_inc(x_452); +x_498 = l_Lean_Syntax_node7(x_452, x_446, x_448, x_448, x_497, x_448, x_448, x_448, x_448); x_499 = 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__4_spec__4___closed__25; lean_inc_ref(x_451); x_500 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_499); x_501 = 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__4_spec__4___closed__26; -lean_inc(x_454); +lean_inc(x_452); x_502 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_502, 0, x_454); +lean_ctor_set(x_502, 0, x_452); lean_ctor_set(x_502, 1, x_501); x_503 = l_Lake_configDecl___closed__8; lean_inc_ref(x_451); x_504 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_503); -lean_inc(x_450); -lean_inc(x_452); -lean_inc(x_504); +lean_inc(x_448); lean_inc(x_454); -x_505 = l_Lean_Syntax_node2(x_454, x_504, x_452, x_450); +lean_inc(x_504); +lean_inc(x_452); +x_505 = l_Lean_Syntax_node2(x_452, x_504, x_454, x_448); x_506 = 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__4_spec__4___closed__27; lean_inc_ref(x_451); x_507 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_506); -lean_inc_n(x_450, 2); -lean_inc(x_454); -x_508 = l_Lean_Syntax_node2(x_454, x_507, x_450, x_450); +lean_inc_n(x_448, 2); +lean_inc(x_452); +x_508 = l_Lean_Syntax_node2(x_452, x_507, x_448, x_448); x_509 = 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__2_spec__2___closed__18; lean_inc_ref(x_451); x_510 = l_Lean_Name_mkStr4(x_19, x_20, x_451, x_509); x_511 = 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__2_spec__2___closed__19; -lean_inc(x_454); +lean_inc(x_452); x_512 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_512, 0, x_454); +lean_ctor_set(x_512, 0, x_452); lean_ctor_set(x_512, 1, x_511); x_513 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___closed__56; -lean_inc_n(x_450, 2); -lean_inc(x_454); -x_514 = l_Lean_Syntax_node2(x_454, x_513, x_450, x_450); -lean_inc(x_450); +lean_inc_n(x_448, 2); +lean_inc(x_452); +x_514 = l_Lean_Syntax_node2(x_452, x_513, x_448, x_448); +lean_inc(x_448); lean_inc(x_510); -lean_inc(x_454); -x_515 = l_Lean_Syntax_node4(x_454, x_510, x_512, x_495, x_514, x_450); -lean_inc(x_454); -x_516 = l_Lean_Syntax_node5(x_454, x_500, x_502, x_505, x_508, x_515, x_450); -lean_inc(x_447); -x_517 = l_Lean_Syntax_node2(x_454, x_447, x_498, x_516); -x_518 = l_Lean_Name_hasMacroScopes(x_448); +lean_inc(x_452); +x_515 = l_Lean_Syntax_node4(x_452, x_510, x_512, x_495, x_514, x_448); +lean_inc(x_452); +x_516 = l_Lean_Syntax_node5(x_452, x_500, x_502, x_505, x_508, x_515, x_448); +lean_inc(x_453); +x_517 = l_Lean_Syntax_node2(x_452, x_453, x_498, x_516); +x_518 = l_Lean_Name_hasMacroScopes(x_447); if (x_518 == 0) { lean_object* x_519; -lean_inc(x_448); -x_519 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__3(x_448); -x_424 = x_446; -x_425 = x_447; -x_426 = x_510; -x_427 = x_513; -x_428 = x_448; -x_429 = x_511; -x_430 = x_451; -x_431 = x_517; -x_432 = x_452; -x_433 = x_504; -x_434 = x_494; -x_435 = x_453; +lean_inc(x_447); +x_519 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__3(x_447); +x_424 = x_510; +x_425 = x_511; +x_426 = x_517; +x_427 = x_446; +x_428 = x_447; +x_429 = x_513; +x_430 = x_504; +x_431 = x_451; +x_432 = x_450; +x_433 = x_453; +x_434 = x_454; +x_435 = x_494; x_436 = x_519; goto block_445; } else { 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_object* x_526; lean_object* x_527; lean_object* x_528; -lean_inc(x_448); -x_520 = l_Lean_extractMacroScopes(x_448); +lean_inc(x_447); +x_520 = l_Lean_extractMacroScopes(x_447); x_521 = lean_ctor_get(x_520, 0); lean_inc(x_521); x_522 = lean_ctor_get(x_520, 1); @@ -13450,18 +13450,18 @@ lean_ctor_set(x_527, 1, x_522); lean_ctor_set(x_527, 2, x_523); lean_ctor_set(x_527, 3, x_524); x_528 = l_Lean_MacroScopesView_review(x_527); -x_424 = x_446; -x_425 = x_447; -x_426 = x_510; -x_427 = x_513; -x_428 = x_448; -x_429 = x_511; -x_430 = x_451; -x_431 = x_517; -x_432 = x_452; -x_433 = x_504; -x_434 = x_494; -x_435 = x_453; +x_424 = x_510; +x_425 = x_511; +x_426 = x_517; +x_427 = x_446; +x_428 = x_447; +x_429 = x_513; +x_430 = x_504; +x_431 = x_451; +x_432 = x_450; +x_433 = x_453; +x_434 = x_454; +x_435 = x_494; x_436 = x_528; goto block_445; } @@ -13470,7 +13470,7 @@ goto block_445; block_545: { 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; -x_534 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_530); +x_534 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__0(x_10, x_6, x_531); x_535 = lean_ctor_get(x_534, 0); lean_inc(x_535); x_536 = lean_ctor_get(x_534, 1); @@ -13490,14 +13490,14 @@ if (lean_obj_tag(x_1) == 0) lean_object* x_542; x_542 = 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__2_spec__2___closed__53; x_446 = x_540; -x_447 = x_539; -x_448 = x_532; -x_449 = x_531; -x_450 = x_541; +x_447 = x_530; +x_448 = x_541; +x_449 = x_532; +x_450 = x_536; x_451 = x_538; -x_452 = x_537; -x_453 = x_536; -x_454 = x_535; +x_452 = x_535; +x_453 = x_539; +x_454 = x_537; x_455 = x_542; goto block_529; } @@ -13508,14 +13508,14 @@ x_543 = lean_ctor_get(x_1, 0); lean_inc(x_543); x_544 = l_Array_mkArray1___redArg(x_543); x_446 = x_540; -x_447 = x_539; -x_448 = x_532; -x_449 = x_531; -x_450 = x_541; +x_447 = x_530; +x_448 = x_541; +x_449 = x_532; +x_450 = x_536; x_451 = x_538; -x_452 = x_537; -x_453 = x_536; -x_454 = x_535; +x_452 = x_535; +x_453 = x_539; +x_454 = x_537; x_455 = x_544; goto block_529; } @@ -13530,9 +13530,9 @@ if (x_549 == 0) lean_object* x_550; lean_inc(x_548); x_550 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__4(x_548); -x_530 = x_547; -x_531 = x_546; -x_532 = x_548; +x_530 = x_548; +x_531 = x_547; +x_532 = x_546; x_533 = x_550; goto block_545; } @@ -13549,9 +13549,9 @@ x_553 = lean_ctor_get(x_551, 0); x_554 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls___lam__4(x_553); lean_ctor_set(x_551, 0, x_554); x_555 = l_Lean_MacroScopesView_review(x_551); -x_530 = x_547; -x_531 = x_546; -x_532 = x_548; +x_530 = x_548; +x_531 = x_547; +x_532 = x_546; x_533 = x_555; goto block_545; } @@ -13574,9 +13574,9 @@ lean_ctor_set(x_561, 1, x_557); lean_ctor_set(x_561, 2, x_558); lean_ctor_set(x_561, 3, x_559); x_562 = l_Lean_MacroScopesView_review(x_561); -x_530 = x_547; -x_531 = x_546; -x_532 = x_548; +x_530 = x_548; +x_531 = x_547; +x_532 = x_546; x_533 = x_562; goto block_545; } @@ -15242,8 +15242,8 @@ x_29 = l___private_Lake_Config_Meta_0__Lake_mkParentFieldView___lam__0(x_27); lean_dec(x_27); x_4 = x_23; x_5 = x_24; -x_6 = x_25; -x_7 = x_26; +x_6 = x_26; +x_7 = x_25; x_8 = x_29; goto block_12; } @@ -15262,8 +15262,8 @@ lean_ctor_set(x_30, 0, x_33); x_34 = l_Lean_MacroScopesView_review(x_30); x_4 = x_23; x_5 = x_24; -x_6 = x_25; -x_7 = x_26; +x_6 = x_26; +x_7 = x_25; x_8 = x_34; goto block_12; } @@ -15289,8 +15289,8 @@ lean_ctor_set(x_40, 3, x_38); x_41 = l_Lean_MacroScopesView_review(x_40); x_4 = x_23; x_5 = x_24; -x_6 = x_25; -x_7 = x_26; +x_6 = x_26; +x_7 = x_25; x_8 = x_41; goto block_12; } @@ -15555,8 +15555,8 @@ x_106 = l___private_Lake_Config_Meta_0__Lake_mkParentFieldView___lam__0(x_104); lean_dec(x_104); x_4 = x_100; x_5 = x_101; -x_6 = x_102; -x_7 = x_103; +x_6 = x_103; +x_7 = x_102; x_8 = x_106; goto block_12; } @@ -15596,8 +15596,8 @@ lean_ctor_set(x_114, 3, x_111); x_115 = l_Lean_MacroScopesView_review(x_114); x_4 = x_100; x_5 = x_101; -x_6 = x_102; -x_7 = x_103; +x_6 = x_103; +x_7 = x_102; x_8 = x_115; goto block_12; } @@ -15774,7 +15774,7 @@ uint8_t x_9; lean_object* x_10; lean_object* x_11; x_9 = 0; x_10 = l_Lean_mkIdentFrom(x_5, x_8, x_9); lean_dec(x_5); -x_11 = lean_apply_3(x_4, x_10, x_6, x_7); +x_11 = lean_apply_3(x_4, x_10, x_7, x_6); return x_11; } } @@ -16216,14 +16216,14 @@ goto block_230; lean_object* x_40; lean_object* x_41; uint8_t x_42; x_40 = lean_array_get_size(x_23); lean_dec_ref(x_23); -x_41 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_39, x_30, x_40, x_34, x_35, x_19, x_38); +x_41 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_39, x_30, x_40, x_34, x_38, x_19, x_36); x_42 = !lean_is_exclusive(x_41); if (x_42 == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_43 = lean_ctor_get(x_41, 0); x_44 = lean_mk_empty_array_with_capacity(x_6); -x_45 = lean_array_push(x_44, x_36); +x_45 = lean_array_push(x_44, x_35); x_46 = l_Array_append___redArg(x_45, x_43); lean_dec(x_43); x_47 = lean_box(2); @@ -16243,7 +16243,7 @@ lean_inc(x_50); lean_inc(x_49); lean_dec(x_41); x_51 = lean_mk_empty_array_with_capacity(x_6); -x_52 = lean_array_push(x_51, x_36); +x_52 = lean_array_push(x_51, x_35); x_53 = l_Array_append___redArg(x_52, x_49); lean_dec(x_49); x_54 = lean_box(2); @@ -16260,40 +16260,40 @@ return x_56; block_93: { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; size_t x_75; lean_object* x_76; size_t x_77; lean_object* x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_inc_ref(x_59); -x_71 = l_Array_append___redArg(x_59, x_70); +lean_inc_ref(x_69); +x_71 = l_Array_append___redArg(x_69, x_70); lean_dec_ref(x_70); lean_inc(x_65); -lean_inc(x_69); +lean_inc(x_59); x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 0, x_59); lean_ctor_set(x_72, 1, x_65); lean_ctor_set(x_72, 2, x_71); x_73 = l_Lake_expandConfigDecl___lam__1___closed__0; x_74 = l_Lean_Name_mkStr4(x_7, x_8, x_9, x_73); -x_75 = lean_array_size(x_64); -x_76 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_75, x_32, x_64); +x_75 = lean_array_size(x_68); +x_76 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_75, x_32, x_68); x_77 = lean_array_size(x_76); x_78 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_77, x_32, x_76); x_79 = lean_array_size(x_78); x_80 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_79, x_32, x_78); -x_81 = l_Array_append___redArg(x_59, x_80); +x_81 = l_Array_append___redArg(x_69, x_80); lean_dec_ref(x_80); lean_inc(x_65); -lean_inc(x_69); +lean_inc(x_59); x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_69); +lean_ctor_set(x_82, 0, x_59); lean_ctor_set(x_82, 1, x_65); lean_ctor_set(x_82, 2, x_81); -lean_inc(x_69); -x_83 = l_Lean_Syntax_node1(x_69, x_74, x_82); +lean_inc(x_59); +x_83 = l_Lean_Syntax_node1(x_59, x_74, x_82); lean_inc(x_65); -lean_inc(x_69); -x_84 = l_Lean_Syntax_node3(x_69, x_65, x_60, x_72, x_83); -lean_inc(x_69); -x_85 = l_Lean_Syntax_node6(x_69, x_58, x_63, x_4, x_68, x_62, x_84, x_29); +lean_inc(x_59); +x_84 = l_Lean_Syntax_node3(x_59, x_65, x_64, x_72, x_83); +lean_inc(x_59); +x_85 = l_Lean_Syntax_node6(x_59, x_62, x_67, x_4, x_60, x_66, x_84, x_29); lean_inc(x_10); -x_86 = l_Lean_Syntax_node2(x_69, x_67, x_10, x_85); +x_86 = l_Lean_Syntax_node2(x_59, x_58, x_10, x_85); x_87 = l_Lean_Syntax_getArg(x_10, x_11); lean_dec(x_10); x_88 = l_Lean_Syntax_getOptional_x3f(x_87); @@ -16302,10 +16302,10 @@ if (lean_obj_tag(x_88) == 0) { lean_object* x_89; x_89 = lean_box(0); -x_35 = x_61; -x_36 = x_86; +x_35 = x_86; +x_36 = x_61; x_37 = x_65; -x_38 = x_66; +x_38 = x_63; x_39 = x_89; goto block_57; } @@ -16315,10 +16315,10 @@ uint8_t x_90; x_90 = !lean_is_exclusive(x_88); if (x_90 == 0) { -x_35 = x_61; -x_36 = x_86; +x_35 = x_86; +x_36 = x_61; x_37 = x_65; -x_38 = x_66; +x_38 = x_63; x_39 = x_88; goto block_57; } @@ -16330,10 +16330,10 @@ lean_inc(x_91); lean_dec(x_88); x_92 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_92, 0, x_91); -x_35 = x_61; -x_36 = x_86; +x_35 = x_86; +x_36 = x_61; x_37 = x_65; -x_38 = x_66; +x_38 = x_63; x_39 = x_92; goto block_57; } @@ -16343,57 +16343,57 @@ goto block_57; { lean_object* x_106; x_106 = 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__2_spec__2___closed__53; -x_58 = x_95; -x_59 = x_94; +x_58 = x_94; +x_59 = x_95; x_60 = x_96; x_61 = x_97; -x_62 = x_100; -x_63 = x_99; -x_64 = x_98; -x_65 = x_101; -x_66 = x_103; -x_67 = x_102; -x_68 = x_105; -x_69 = x_104; +x_62 = x_98; +x_63 = x_101; +x_64 = x_100; +x_65 = x_99; +x_66 = x_102; +x_67 = x_105; +x_68 = x_104; +x_69 = x_103; x_70 = x_106; goto block_93; } block_131: { 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_inc_ref(x_108); -x_122 = l_Array_append___redArg(x_108, x_121); +lean_inc_ref(x_119); +x_122 = l_Array_append___redArg(x_119, x_121); lean_dec_ref(x_121); -lean_inc(x_111); -lean_inc(x_119); +lean_inc(x_117); +lean_inc(x_115); x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_119); -lean_ctor_set(x_123, 1, x_111); +lean_ctor_set(x_123, 0, x_115); +lean_ctor_set(x_123, 1, x_117); lean_ctor_set(x_123, 2, x_122); -lean_inc(x_119); -x_124 = l_Lean_Syntax_node3(x_119, x_114, x_120, x_116, x_123); -lean_inc(x_111); -lean_inc(x_119); -x_125 = l_Lean_Syntax_node1(x_119, x_111, x_124); +lean_inc(x_115); +x_124 = l_Lean_Syntax_node3(x_115, x_118, x_108, x_109, x_123); +lean_inc(x_117); +lean_inc(x_115); +x_125 = l_Lean_Syntax_node1(x_115, x_117, x_124); x_126 = 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__4_spec__4___closed__37; -lean_inc(x_119); +lean_inc(x_115); x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_119); +lean_ctor_set(x_127, 0, x_115); lean_ctor_set(x_127, 1, x_126); if (lean_obj_tag(x_17) == 0) { lean_dec_ref(x_12); -x_94 = x_108; -x_95 = x_109; -x_96 = x_127; -x_97 = x_115; -x_98 = x_117; -x_99 = x_110; -x_100 = x_125; -x_101 = x_111; -x_102 = x_112; -x_103 = x_118; -x_104 = x_119; +x_94 = x_114; +x_95 = x_115; +x_96 = x_110; +x_97 = x_111; +x_98 = x_116; +x_99 = x_117; +x_100 = x_127; +x_101 = x_112; +x_102 = x_125; +x_103 = x_119; +x_104 = x_120; x_105 = x_113; goto block_107; } @@ -16406,17 +16406,17 @@ lean_dec_ref(x_17); if (lean_obj_tag(x_128) == 0) { lean_dec_ref(x_12); -x_94 = x_108; -x_95 = x_109; -x_96 = x_127; -x_97 = x_115; -x_98 = x_117; -x_99 = x_110; -x_100 = x_125; -x_101 = x_111; -x_102 = x_112; -x_103 = x_118; -x_104 = x_119; +x_94 = x_114; +x_95 = x_115; +x_96 = x_110; +x_97 = x_111; +x_98 = x_116; +x_99 = x_117; +x_100 = x_127; +x_101 = x_112; +x_102 = x_125; +x_103 = x_119; +x_104 = x_120; x_105 = x_113; goto block_107; } @@ -16427,17 +16427,17 @@ x_129 = lean_ctor_get(x_128, 0); lean_inc(x_129); lean_dec_ref(x_128); x_130 = lean_apply_1(x_12, x_129); -x_58 = x_109; -x_59 = x_108; -x_60 = x_127; -x_61 = x_115; -x_62 = x_125; -x_63 = x_110; -x_64 = x_117; -x_65 = x_111; -x_66 = x_118; -x_67 = x_112; -x_68 = x_113; +x_58 = x_114; +x_59 = x_115; +x_60 = x_110; +x_61 = x_111; +x_62 = x_116; +x_63 = x_112; +x_64 = x_127; +x_65 = x_117; +x_66 = x_125; +x_67 = x_113; +x_68 = x_120; x_69 = x_119; x_70 = x_130; goto block_93; @@ -16467,50 +16467,50 @@ goto block_131; block_171: { 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_inc_ref(x_148); -x_160 = l_Array_append___redArg(x_148, x_159); +lean_inc_ref(x_157); +x_160 = l_Array_append___redArg(x_157, x_159); lean_dec_ref(x_159); lean_inc(x_154); -lean_inc(x_158); +lean_inc(x_149); x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 0, x_149); lean_ctor_set(x_161, 1, x_154); lean_ctor_set(x_161, 2, x_160); -lean_inc(x_158); -x_162 = l_Lean_Syntax_node2(x_158, x_149, x_155, x_161); +lean_inc(x_149); +x_162 = l_Lean_Syntax_node2(x_149, x_152, x_158, x_161); x_163 = l_Lake_configDecl___closed__32; lean_inc_ref(x_9); lean_inc_ref(x_8); lean_inc_ref(x_7); x_164 = l_Lean_Name_mkStr4(x_7, x_8, x_9, x_163); -lean_inc(x_158); +lean_inc(x_149); x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_158); +lean_ctor_set(x_165, 0, x_149); lean_ctor_set(x_165, 1, x_163); -lean_inc_ref(x_148); -x_166 = l_Array_append___redArg(x_148, x_153); -lean_dec_ref(x_153); +lean_inc_ref(x_157); +x_166 = l_Array_append___redArg(x_157, x_148); +lean_dec_ref(x_148); lean_inc(x_154); -lean_inc(x_158); +lean_inc(x_149); x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_158); +lean_ctor_set(x_167, 0, x_149); lean_ctor_set(x_167, 1, x_154); lean_ctor_set(x_167, 2, x_166); if (lean_obj_tag(x_13) == 0) { -x_132 = x_148; -x_133 = x_147; -x_134 = x_152; -x_135 = x_154; -x_136 = x_157; -x_137 = x_162; -x_138 = x_164; -x_139 = x_150; -x_140 = x_167; -x_141 = x_151; -x_142 = x_156; -x_143 = x_158; -x_144 = x_165; +x_132 = x_165; +x_133 = x_167; +x_134 = x_162; +x_135 = x_150; +x_136 = x_153; +x_137 = x_156; +x_138 = x_147; +x_139 = x_149; +x_140 = x_151; +x_141 = x_154; +x_142 = x_164; +x_143 = x_157; +x_144 = x_155; goto block_146; } else @@ -16521,19 +16521,19 @@ lean_inc(x_168); lean_dec_ref(x_13); if (lean_obj_tag(x_168) == 0) { -x_132 = x_148; -x_133 = x_147; -x_134 = x_152; -x_135 = x_154; -x_136 = x_157; -x_137 = x_162; -x_138 = x_164; -x_139 = x_150; -x_140 = x_167; -x_141 = x_151; -x_142 = x_156; -x_143 = x_158; -x_144 = x_165; +x_132 = x_165; +x_133 = x_167; +x_134 = x_162; +x_135 = x_150; +x_136 = x_153; +x_137 = x_156; +x_138 = x_147; +x_139 = x_149; +x_140 = x_151; +x_141 = x_154; +x_142 = x_164; +x_143 = x_157; +x_144 = x_155; goto block_146; } else @@ -16544,19 +16544,19 @@ lean_inc(x_169); lean_dec_ref(x_168); lean_inc_ref(x_12); x_170 = lean_apply_1(x_12, x_169); -x_108 = x_148; -x_109 = x_147; -x_110 = x_152; -x_111 = x_154; -x_112 = x_157; -x_113 = x_162; -x_114 = x_164; -x_115 = x_150; -x_116 = x_167; -x_117 = x_151; -x_118 = x_156; -x_119 = x_158; -x_120 = x_165; +x_108 = x_165; +x_109 = x_167; +x_110 = x_162; +x_111 = x_150; +x_112 = x_153; +x_113 = x_156; +x_114 = x_147; +x_115 = x_149; +x_116 = x_151; +x_117 = x_154; +x_118 = x_164; +x_119 = x_157; +x_120 = x_155; x_121 = x_170; goto block_131; } @@ -16615,18 +16615,18 @@ if (lean_obj_tag(x_14) == 0) { lean_object* x_197; x_197 = 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__2_spec__2___closed__53; -x_147 = x_182; -x_148 = x_190; -x_149 = x_188; -x_150 = x_173; -x_151 = x_176; -x_152 = x_186; -x_153 = x_172; +x_147 = x_180; +x_148 = x_172; +x_149 = x_178; +x_150 = x_174; +x_151 = x_182; +x_152 = x_188; +x_153 = x_173; x_154 = x_189; -x_155 = x_196; -x_156 = x_174; -x_157 = x_180; -x_158 = x_178; +x_155 = x_176; +x_156 = x_186; +x_157 = x_190; +x_158 = x_196; x_159 = x_197; goto block_171; } @@ -16637,18 +16637,18 @@ x_198 = lean_ctor_get(x_14, 0); lean_inc(x_198); lean_dec_ref(x_14); x_199 = l_Array_mkArray1___redArg(x_198); -x_147 = x_182; -x_148 = x_190; -x_149 = x_188; -x_150 = x_173; -x_151 = x_176; -x_152 = x_186; -x_153 = x_172; +x_147 = x_180; +x_148 = x_172; +x_149 = x_178; +x_150 = x_174; +x_151 = x_182; +x_152 = x_188; +x_153 = x_173; x_154 = x_189; -x_155 = x_196; -x_156 = x_174; -x_157 = x_180; -x_158 = x_178; +x_155 = x_176; +x_156 = x_186; +x_157 = x_190; +x_158 = x_196; x_159 = x_199; goto block_171; } @@ -16925,7 +16925,7 @@ goto block_445; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; x_261 = lean_array_get_size(x_243); lean_dec_ref(x_243); -x_262 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_260, x_251, x_261, x_255, x_256, x_245, x_259); +x_262 = l___private_Lake_Config_Meta_0__Lake_mkConfigAuxDecls(x_260, x_251, x_261, x_255, x_259, x_245, x_257); x_263 = lean_ctor_get(x_262, 0); lean_inc(x_263); x_264 = lean_ctor_get(x_262, 1); @@ -16939,7 +16939,7 @@ if (lean_is_exclusive(x_262)) { x_265 = lean_box(0); } x_266 = lean_mk_empty_array_with_capacity(x_6); -x_267 = lean_array_push(x_266, x_257); +x_267 = lean_array_push(x_266, x_256); x_268 = l_Array_append___redArg(x_267, x_263); lean_dec(x_263); x_269 = lean_box(2); @@ -16959,40 +16959,40 @@ return x_271; block_308: { lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; size_t x_290; lean_object* x_291; size_t x_292; lean_object* x_293; size_t x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; -lean_inc_ref(x_274); -x_286 = l_Array_append___redArg(x_274, x_285); +lean_inc_ref(x_284); +x_286 = l_Array_append___redArg(x_284, x_285); lean_dec_ref(x_285); lean_inc(x_280); -lean_inc(x_284); +lean_inc(x_274); x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_284); +lean_ctor_set(x_287, 0, x_274); lean_ctor_set(x_287, 1, x_280); lean_ctor_set(x_287, 2, x_286); x_288 = l_Lake_expandConfigDecl___lam__1___closed__0; x_289 = l_Lean_Name_mkStr4(x_7, x_8, x_9, x_288); -x_290 = lean_array_size(x_279); -x_291 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_290, x_253, x_279); +x_290 = lean_array_size(x_283); +x_291 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_290, x_253, x_283); x_292 = lean_array_size(x_291); x_293 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_292, x_253, x_291); x_294 = lean_array_size(x_293); x_295 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake_expandConfigDecl_spec__6___redArg(x_294, x_253, x_293); -x_296 = l_Array_append___redArg(x_274, x_295); +x_296 = l_Array_append___redArg(x_284, x_295); lean_dec_ref(x_295); lean_inc(x_280); -lean_inc(x_284); +lean_inc(x_274); x_297 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_297, 0, x_284); +lean_ctor_set(x_297, 0, x_274); lean_ctor_set(x_297, 1, x_280); lean_ctor_set(x_297, 2, x_296); -lean_inc(x_284); -x_298 = l_Lean_Syntax_node1(x_284, x_289, x_297); +lean_inc(x_274); +x_298 = l_Lean_Syntax_node1(x_274, x_289, x_297); lean_inc(x_280); -lean_inc(x_284); -x_299 = l_Lean_Syntax_node3(x_284, x_280, x_275, x_287, x_298); -lean_inc(x_284); -x_300 = l_Lean_Syntax_node6(x_284, x_273, x_278, x_4, x_283, x_277, x_299, x_250); +lean_inc(x_274); +x_299 = l_Lean_Syntax_node3(x_274, x_280, x_279, x_287, x_298); +lean_inc(x_274); +x_300 = l_Lean_Syntax_node6(x_274, x_277, x_282, x_4, x_275, x_281, x_299, x_250); lean_inc(x_10); -x_301 = l_Lean_Syntax_node2(x_284, x_282, x_10, x_300); +x_301 = l_Lean_Syntax_node2(x_274, x_273, x_10, x_300); x_302 = l_Lean_Syntax_getArg(x_10, x_11); lean_dec(x_10); x_303 = l_Lean_Syntax_getOptional_x3f(x_302); @@ -17001,10 +17001,10 @@ if (lean_obj_tag(x_303) == 0) { lean_object* x_304; x_304 = lean_box(0); -x_256 = x_276; -x_257 = x_301; +x_256 = x_301; +x_257 = x_276; x_258 = x_280; -x_259 = x_281; +x_259 = x_278; x_260 = x_304; goto block_272; } @@ -17026,10 +17026,10 @@ if (lean_is_scalar(x_306)) { x_307 = x_306; } lean_ctor_set(x_307, 0, x_305); -x_256 = x_276; -x_257 = x_301; +x_256 = x_301; +x_257 = x_276; x_258 = x_280; -x_259 = x_281; +x_259 = x_278; x_260 = x_307; goto block_272; } @@ -17038,57 +17038,57 @@ goto block_272; { lean_object* x_321; x_321 = 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__2_spec__2___closed__53; -x_273 = x_310; -x_274 = x_309; +x_273 = x_309; +x_274 = x_310; x_275 = x_311; x_276 = x_312; -x_277 = x_315; -x_278 = x_314; -x_279 = x_313; -x_280 = x_316; -x_281 = x_318; -x_282 = x_317; -x_283 = x_320; -x_284 = x_319; +x_277 = x_313; +x_278 = x_316; +x_279 = x_315; +x_280 = x_314; +x_281 = x_317; +x_282 = x_320; +x_283 = x_319; +x_284 = x_318; x_285 = x_321; goto block_308; } block_346: { lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -lean_inc_ref(x_323); -x_337 = l_Array_append___redArg(x_323, x_336); +lean_inc_ref(x_334); +x_337 = l_Array_append___redArg(x_334, x_336); lean_dec_ref(x_336); -lean_inc(x_326); -lean_inc(x_334); +lean_inc(x_332); +lean_inc(x_330); x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_334); -lean_ctor_set(x_338, 1, x_326); +lean_ctor_set(x_338, 0, x_330); +lean_ctor_set(x_338, 1, x_332); lean_ctor_set(x_338, 2, x_337); -lean_inc(x_334); -x_339 = l_Lean_Syntax_node3(x_334, x_329, x_335, x_331, x_338); -lean_inc(x_326); -lean_inc(x_334); -x_340 = l_Lean_Syntax_node1(x_334, x_326, x_339); +lean_inc(x_330); +x_339 = l_Lean_Syntax_node3(x_330, x_333, x_323, x_324, x_338); +lean_inc(x_332); +lean_inc(x_330); +x_340 = l_Lean_Syntax_node1(x_330, x_332, x_339); x_341 = 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__4_spec__4___closed__37; -lean_inc(x_334); +lean_inc(x_330); x_342 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_342, 0, x_334); +lean_ctor_set(x_342, 0, x_330); lean_ctor_set(x_342, 1, x_341); if (lean_obj_tag(x_17) == 0) { lean_dec_ref(x_12); -x_309 = x_323; -x_310 = x_324; -x_311 = x_342; -x_312 = x_330; -x_313 = x_332; -x_314 = x_325; -x_315 = x_340; -x_316 = x_326; -x_317 = x_327; -x_318 = x_333; -x_319 = x_334; +x_309 = x_329; +x_310 = x_330; +x_311 = x_325; +x_312 = x_326; +x_313 = x_331; +x_314 = x_332; +x_315 = x_342; +x_316 = x_327; +x_317 = x_340; +x_318 = x_334; +x_319 = x_335; x_320 = x_328; goto block_322; } @@ -17101,17 +17101,17 @@ lean_dec_ref(x_17); if (lean_obj_tag(x_343) == 0) { lean_dec_ref(x_12); -x_309 = x_323; -x_310 = x_324; -x_311 = x_342; -x_312 = x_330; -x_313 = x_332; -x_314 = x_325; -x_315 = x_340; -x_316 = x_326; -x_317 = x_327; -x_318 = x_333; -x_319 = x_334; +x_309 = x_329; +x_310 = x_330; +x_311 = x_325; +x_312 = x_326; +x_313 = x_331; +x_314 = x_332; +x_315 = x_342; +x_316 = x_327; +x_317 = x_340; +x_318 = x_334; +x_319 = x_335; x_320 = x_328; goto block_322; } @@ -17122,17 +17122,17 @@ x_344 = lean_ctor_get(x_343, 0); lean_inc(x_344); lean_dec_ref(x_343); x_345 = lean_apply_1(x_12, x_344); -x_273 = x_324; -x_274 = x_323; -x_275 = x_342; -x_276 = x_330; -x_277 = x_340; -x_278 = x_325; -x_279 = x_332; -x_280 = x_326; -x_281 = x_333; -x_282 = x_327; -x_283 = x_328; +x_273 = x_329; +x_274 = x_330; +x_275 = x_325; +x_276 = x_326; +x_277 = x_331; +x_278 = x_327; +x_279 = x_342; +x_280 = x_332; +x_281 = x_340; +x_282 = x_328; +x_283 = x_335; x_284 = x_334; x_285 = x_345; goto block_308; @@ -17162,50 +17162,50 @@ goto block_346; block_386: { 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_inc_ref(x_363); -x_375 = l_Array_append___redArg(x_363, x_374); +lean_inc_ref(x_372); +x_375 = l_Array_append___redArg(x_372, x_374); lean_dec_ref(x_374); lean_inc(x_369); -lean_inc(x_373); +lean_inc(x_364); x_376 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_376, 0, x_373); +lean_ctor_set(x_376, 0, x_364); lean_ctor_set(x_376, 1, x_369); lean_ctor_set(x_376, 2, x_375); -lean_inc(x_373); -x_377 = l_Lean_Syntax_node2(x_373, x_364, x_370, x_376); +lean_inc(x_364); +x_377 = l_Lean_Syntax_node2(x_364, x_367, x_373, x_376); x_378 = l_Lake_configDecl___closed__32; lean_inc_ref(x_9); lean_inc_ref(x_8); lean_inc_ref(x_7); x_379 = l_Lean_Name_mkStr4(x_7, x_8, x_9, x_378); -lean_inc(x_373); +lean_inc(x_364); x_380 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_380, 0, x_373); +lean_ctor_set(x_380, 0, x_364); lean_ctor_set(x_380, 1, x_378); -lean_inc_ref(x_363); -x_381 = l_Array_append___redArg(x_363, x_368); -lean_dec_ref(x_368); +lean_inc_ref(x_372); +x_381 = l_Array_append___redArg(x_372, x_363); +lean_dec_ref(x_363); lean_inc(x_369); -lean_inc(x_373); +lean_inc(x_364); x_382 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_382, 0, x_373); +lean_ctor_set(x_382, 0, x_364); lean_ctor_set(x_382, 1, x_369); lean_ctor_set(x_382, 2, x_381); if (lean_obj_tag(x_13) == 0) { -x_347 = x_363; -x_348 = x_362; -x_349 = x_367; -x_350 = x_369; -x_351 = x_372; -x_352 = x_377; -x_353 = x_379; -x_354 = x_365; -x_355 = x_382; -x_356 = x_366; -x_357 = x_371; -x_358 = x_373; -x_359 = x_380; +x_347 = x_380; +x_348 = x_382; +x_349 = x_377; +x_350 = x_365; +x_351 = x_368; +x_352 = x_371; +x_353 = x_362; +x_354 = x_364; +x_355 = x_366; +x_356 = x_369; +x_357 = x_379; +x_358 = x_372; +x_359 = x_370; goto block_361; } else @@ -17216,19 +17216,19 @@ lean_inc(x_383); lean_dec_ref(x_13); if (lean_obj_tag(x_383) == 0) { -x_347 = x_363; -x_348 = x_362; -x_349 = x_367; -x_350 = x_369; -x_351 = x_372; -x_352 = x_377; -x_353 = x_379; -x_354 = x_365; -x_355 = x_382; -x_356 = x_366; -x_357 = x_371; -x_358 = x_373; -x_359 = x_380; +x_347 = x_380; +x_348 = x_382; +x_349 = x_377; +x_350 = x_365; +x_351 = x_368; +x_352 = x_371; +x_353 = x_362; +x_354 = x_364; +x_355 = x_366; +x_356 = x_369; +x_357 = x_379; +x_358 = x_372; +x_359 = x_370; goto block_361; } else @@ -17239,19 +17239,19 @@ lean_inc(x_384); lean_dec_ref(x_383); lean_inc_ref(x_12); x_385 = lean_apply_1(x_12, x_384); -x_323 = x_363; -x_324 = x_362; -x_325 = x_367; -x_326 = x_369; -x_327 = x_372; -x_328 = x_377; -x_329 = x_379; -x_330 = x_365; -x_331 = x_382; -x_332 = x_366; -x_333 = x_371; -x_334 = x_373; -x_335 = x_380; +x_323 = x_380; +x_324 = x_382; +x_325 = x_377; +x_326 = x_365; +x_327 = x_368; +x_328 = x_371; +x_329 = x_362; +x_330 = x_364; +x_331 = x_366; +x_332 = x_369; +x_333 = x_379; +x_334 = x_372; +x_335 = x_370; x_336 = x_385; goto block_346; } @@ -17310,18 +17310,18 @@ if (lean_obj_tag(x_14) == 0) { lean_object* x_412; x_412 = 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__2_spec__2___closed__53; -x_362 = x_397; -x_363 = x_405; -x_364 = x_403; -x_365 = x_388; -x_366 = x_391; -x_367 = x_401; -x_368 = x_387; +x_362 = x_395; +x_363 = x_387; +x_364 = x_393; +x_365 = x_389; +x_366 = x_397; +x_367 = x_403; +x_368 = x_388; x_369 = x_404; -x_370 = x_411; -x_371 = x_389; -x_372 = x_395; -x_373 = x_393; +x_370 = x_391; +x_371 = x_401; +x_372 = x_405; +x_373 = x_411; x_374 = x_412; goto block_386; } @@ -17332,18 +17332,18 @@ x_413 = lean_ctor_get(x_14, 0); lean_inc(x_413); lean_dec_ref(x_14); x_414 = l_Array_mkArray1___redArg(x_413); -x_362 = x_397; -x_363 = x_405; -x_364 = x_403; -x_365 = x_388; -x_366 = x_391; -x_367 = x_401; -x_368 = x_387; +x_362 = x_395; +x_363 = x_387; +x_364 = x_393; +x_365 = x_389; +x_366 = x_397; +x_367 = x_403; +x_368 = x_388; x_369 = x_404; -x_370 = x_411; -x_371 = x_389; -x_372 = x_395; -x_373 = x_393; +x_370 = x_391; +x_371 = x_401; +x_372 = x_405; +x_373 = x_411; x_374 = x_414; goto block_386; } diff --git a/stage0/stdlib/Lake/DSL/Package.c b/stage0/stdlib/Lake/DSL/Package.c index cf591c81f178..da091d647ec7 100644 --- a/stage0/stdlib/Lake/DSL/Package.c +++ b/stage0/stdlib/Lake/DSL/Package.c @@ -5001,9 +5001,9 @@ if (x_391 == 0) lean_object* x_392; lean_object* x_393; lean_dec(x_384); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_392 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_393 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_392, x_381, x_382); lean_dec(x_1); @@ -5022,9 +5022,9 @@ lean_object* x_397; lean_object* x_398; lean_dec(x_394); lean_dec(x_384); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_397 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_398 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_397, x_381, x_382); lean_dec(x_1); @@ -5041,9 +5041,9 @@ lean_object* x_401; lean_object* x_402; lean_dec(x_394); lean_dec(x_384); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_401 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_402 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_401, x_381, x_382); lean_dec(x_1); @@ -5060,9 +5060,9 @@ if (x_404 == 0) lean_object* x_405; lean_object* x_406; lean_dec(x_384); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_405 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_406 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_405, x_381, x_382); lean_dec(x_1); @@ -5072,7 +5072,7 @@ else { lean_object* x_407; lean_object* x_408; uint8_t x_409; x_407 = l_Lean_Syntax_getArg(x_384, x_379); -x_408 = l_Lean_Syntax_getArg(x_384, x_378); +x_408 = l_Lean_Syntax_getArg(x_384, x_374); lean_dec(x_384); x_409 = l_Lean_Syntax_isNone(x_408); if (x_409 == 0) @@ -5086,9 +5086,9 @@ lean_object* x_411; lean_object* x_412; lean_dec(x_408); lean_dec(x_407); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_411 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_412 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_411, x_381, x_382); lean_dec(x_1); @@ -5108,9 +5108,9 @@ lean_object* x_416; lean_object* x_417; lean_dec(x_413); lean_dec(x_407); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_416 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_417 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_416, x_381, x_382); lean_dec(x_1); @@ -5122,18 +5122,18 @@ lean_object* x_418; lean_dec(x_1); x_418 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_418, 0, x_413); -x_149 = x_374; -x_150 = x_380; -x_151 = x_407; +x_149 = x_387; +x_150 = x_390; +x_151 = x_389; x_152 = x_376; x_153 = x_377; -x_154 = x_379; -x_155 = x_387; -x_156 = x_389; -x_157 = x_390; +x_154 = x_380; +x_155 = x_378; +x_156 = x_395; +x_157 = x_388; x_158 = x_386; -x_159 = x_395; -x_160 = x_388; +x_159 = x_379; +x_160 = x_407; x_161 = x_418; x_162 = x_381; x_163 = x_382; @@ -5147,18 +5147,18 @@ lean_object* x_419; lean_dec(x_408); lean_dec(x_1); x_419 = lean_box(0); -x_149 = x_374; -x_150 = x_380; -x_151 = x_407; +x_149 = x_387; +x_150 = x_390; +x_151 = x_389; x_152 = x_376; x_153 = x_377; -x_154 = x_379; -x_155 = x_387; -x_156 = x_389; -x_157 = x_390; +x_154 = x_380; +x_155 = x_378; +x_156 = x_395; +x_157 = x_388; x_158 = x_386; -x_159 = x_395; -x_160 = x_388; +x_159 = x_379; +x_160 = x_407; x_161 = x_419; x_162 = x_381; x_163 = x_382; @@ -5185,9 +5185,9 @@ lean_object* x_426; lean_object* x_427; lean_dec(x_420); lean_dec(x_384); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_426 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_427 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_426, x_381, x_382); lean_dec(x_1); @@ -5212,9 +5212,9 @@ lean_object* x_432; lean_object* x_433; lean_dec(x_429); lean_dec(x_428); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_432 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_433 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_432, x_381, x_382); lean_dec(x_1); @@ -5234,9 +5234,9 @@ lean_object* x_437; lean_object* x_438; lean_dec(x_434); lean_dec(x_428); lean_dec(x_380); +lean_dec(x_378); lean_dec(x_377); lean_dec(x_376); -lean_dec(x_374); x_437 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__24; x_438 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_437, x_381, x_382); lean_dec(x_1); @@ -5248,15 +5248,15 @@ lean_object* x_439; lean_dec(x_1); x_439 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_439, 0, x_434); -x_350 = x_374; -x_351 = x_380; +x_350 = x_422; +x_351 = x_423; x_352 = x_428; -x_353 = x_423; +x_353 = x_424; x_354 = x_376; x_355 = x_377; -x_356 = x_424; +x_356 = x_380; x_357 = x_421; -x_358 = x_422; +x_358 = x_378; x_359 = x_439; x_360 = x_381; x_361 = x_382; @@ -5270,15 +5270,15 @@ lean_object* x_440; lean_dec(x_429); lean_dec(x_1); x_440 = lean_box(0); -x_350 = x_374; -x_351 = x_380; +x_350 = x_422; +x_351 = x_423; x_352 = x_428; -x_353 = x_423; +x_353 = x_424; x_354 = x_376; x_355 = x_377; -x_356 = x_424; +x_356 = x_380; x_357 = x_421; -x_358 = x_422; +x_358 = x_378; x_359 = x_440; x_360 = x_381; x_361 = x_382; @@ -5319,11 +5319,11 @@ x_455 = l_Lean_Syntax_getArg(x_450, x_373); lean_dec(x_450); x_456 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_456, 0, x_455); -x_374 = x_444; +x_374 = x_449; x_375 = x_447; x_376 = x_448; x_377 = x_442; -x_378 = x_449; +x_378 = x_444; x_379 = x_443; x_380 = x_456; x_381 = x_445; @@ -5336,11 +5336,11 @@ else lean_object* x_457; lean_dec(x_450); x_457 = lean_box(0); -x_374 = x_444; +x_374 = x_449; x_375 = x_447; x_376 = x_448; x_377 = x_442; -x_378 = x_449; +x_378 = x_444; x_379 = x_443; x_380 = x_457; x_381 = x_445; @@ -5401,331 +5401,331 @@ goto block_458; block_26: { 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_Array_append___redArg(x_8, x_19); +x_20 = l_Array_append___redArg(x_4, x_19); lean_dec_ref(x_19); -lean_inc(x_11); +lean_inc(x_18); x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_11); -lean_ctor_set(x_21, 1, x_15); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_16); lean_ctor_set(x_21, 2, x_20); -lean_inc(x_11); -x_22 = l_Lean_Syntax_node4(x_11, x_16, x_7, x_13, x_5, x_21); -lean_inc(x_11); -x_23 = l_Lean_Syntax_node5(x_11, x_18, x_4, x_9, x_17, x_22, x_6); -x_24 = l_Lean_Syntax_node2(x_11, x_14, x_10, x_23); +lean_inc(x_18); +x_22 = l_Lean_Syntax_node4(x_18, x_9, x_14, x_11, x_10, x_21); +lean_inc(x_18); +x_23 = l_Lean_Syntax_node5(x_18, x_13, x_17, x_5, x_12, x_22, x_15); +x_24 = l_Lean_Syntax_node2(x_18, x_8, x_7, 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_12); +lean_ctor_set(x_25, 1, x_6); return x_25; } block_148: { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -lean_inc_ref(x_29); -x_48 = l_Array_append___redArg(x_29, x_47); +lean_inc_ref(x_33); +x_48 = l_Array_append___redArg(x_33, x_47); lean_dec_ref(x_47); -lean_inc(x_39); -lean_inc(x_36); +lean_inc(x_42); +lean_inc(x_45); x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_36); -lean_ctor_set(x_49, 1, x_39); +lean_ctor_set(x_49, 0, x_45); +lean_ctor_set(x_49, 1, x_42); lean_ctor_set(x_49, 2, x_48); x_50 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__15; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_51 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_50); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_51 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_50); x_52 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__16; -lean_inc(x_36); +lean_inc(x_45); x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_36); +lean_ctor_set(x_53, 0, x_45); lean_ctor_set(x_53, 1, x_52); x_54 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__38; -x_55 = l_Lean_Syntax_SepArray_ofElems(x_54, x_31); -lean_dec_ref(x_31); -lean_inc_ref(x_29); -x_56 = l_Array_append___redArg(x_29, x_55); +x_55 = l_Lean_Syntax_SepArray_ofElems(x_54, x_36); +lean_dec_ref(x_36); +lean_inc_ref(x_33); +x_56 = l_Array_append___redArg(x_33, x_55); lean_dec_ref(x_55); -lean_inc(x_39); -lean_inc(x_36); +lean_inc(x_42); +lean_inc(x_45); x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_36); -lean_ctor_set(x_57, 1, x_39); +lean_ctor_set(x_57, 0, x_45); +lean_ctor_set(x_57, 1, x_42); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__17; -lean_inc(x_36); +lean_inc(x_45); x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_36); +lean_ctor_set(x_59, 0, x_45); lean_ctor_set(x_59, 1, x_58); -lean_inc(x_36); -x_60 = l_Lean_Syntax_node3(x_36, x_51, x_53, x_57, x_59); -lean_inc(x_39); -lean_inc(x_36); -x_61 = l_Lean_Syntax_node1(x_36, x_39, x_60); -lean_inc_n(x_28, 5); -lean_inc(x_36); -x_62 = l_Lean_Syntax_node7(x_36, x_27, x_49, x_61, x_28, x_28, x_28, x_28, x_28); +lean_inc(x_45); +x_60 = l_Lean_Syntax_node3(x_45, x_51, x_53, x_57, x_59); +lean_inc(x_42); +lean_inc(x_45); +x_61 = l_Lean_Syntax_node1(x_45, x_42, x_60); +lean_inc_n(x_40, 5); +lean_inc(x_45); +x_62 = l_Lean_Syntax_node7(x_45, x_34, x_49, x_61, x_40, x_40, x_40, x_40, x_40); x_63 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__11; -lean_inc_ref(x_33); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_64 = l_Lean_Name_mkStr4(x_42, x_46, x_33, x_63); +lean_inc_ref(x_28); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_64 = l_Lean_Name_mkStr4(x_27, x_44, x_28, x_63); x_65 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__13; -lean_inc(x_36); +lean_inc(x_45); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_36); +lean_ctor_set(x_66, 0, x_45); lean_ctor_set(x_66, 1, x_65); x_67 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__14; -lean_inc_ref(x_33); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_68 = l_Lean_Name_mkStr4(x_42, x_46, x_33, x_67); +lean_inc_ref(x_28); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_68 = l_Lean_Name_mkStr4(x_27, x_44, x_28, x_67); x_69 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__1; x_70 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__2; -lean_inc(x_34); +lean_inc(x_32); +lean_inc(x_29); +x_71 = l_Lean_addMacroScope(x_29, x_70, x_32); lean_inc(x_30); -x_71 = l_Lean_addMacroScope(x_30, x_70, x_34); -lean_inc(x_44); -lean_inc(x_36); +lean_inc(x_45); x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_36); +lean_ctor_set(x_72, 0, x_45); lean_ctor_set(x_72, 1, x_69); lean_ctor_set(x_72, 2, x_71); -lean_ctor_set(x_72, 3, x_44); -lean_inc(x_28); -lean_inc(x_36); -x_73 = l_Lean_Syntax_node2(x_36, x_68, x_72, x_28); +lean_ctor_set(x_72, 3, x_30); +lean_inc(x_40); +lean_inc(x_45); +x_73 = l_Lean_Syntax_node2(x_45, x_68, x_72, x_40); x_74 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__18; -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_75 = l_Lean_Name_mkStr4(x_42, x_46, x_33, x_74); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_75 = l_Lean_Name_mkStr4(x_27, x_44, x_28, x_74); x_76 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__21; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_77 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_76); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_77 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_76); x_78 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__23; -lean_inc(x_36); +lean_inc(x_45); x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_36); +lean_ctor_set(x_79, 0, x_45); lean_ctor_set(x_79, 1, x_78); x_80 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__4; x_81 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__5; -lean_inc(x_34); -lean_inc(x_30); -x_82 = l_Lean_addMacroScope(x_30, x_81, x_34); +lean_inc(x_32); +lean_inc(x_29); +x_82 = l_Lean_addMacroScope(x_29, x_81, x_32); x_83 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__7; x_84 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__8; -lean_inc(x_44); +lean_inc(x_30); x_85 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_44); +lean_ctor_set(x_85, 1, x_30); x_86 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_86, 0, x_83); lean_ctor_set(x_86, 1, x_85); -lean_inc(x_36); +lean_inc(x_45); x_87 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_87, 0, x_36); +lean_ctor_set(x_87, 0, x_45); lean_ctor_set(x_87, 1, x_80); lean_ctor_set(x_87, 2, x_82); lean_ctor_set(x_87, 3, x_86); -lean_inc(x_36); -x_88 = l_Lean_Syntax_node2(x_36, x_77, x_79, x_87); -lean_inc(x_39); -lean_inc(x_36); -x_89 = l_Lean_Syntax_node1(x_36, x_39, x_88); -lean_inc(x_28); -lean_inc(x_36); -x_90 = l_Lean_Syntax_node2(x_36, x_75, x_28, x_89); +lean_inc(x_45); +x_88 = l_Lean_Syntax_node2(x_45, x_77, x_79, x_87); +lean_inc(x_42); +lean_inc(x_45); +x_89 = l_Lean_Syntax_node1(x_45, x_42, x_88); +lean_inc(x_40); +lean_inc(x_45); +x_90 = l_Lean_Syntax_node2(x_45, x_75, x_40, x_89); x_91 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__37; -lean_inc(x_36); +lean_inc(x_45); x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_36); +lean_ctor_set(x_92, 0, x_45); lean_ctor_set(x_92, 1, x_91); x_93 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__29; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_94 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_93); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_94 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_93); x_95 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__30; -lean_inc(x_36); +lean_inc(x_45); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_36); +lean_ctor_set(x_96, 0, x_45); lean_ctor_set(x_96, 1, x_95); x_97 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___closed__12; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_98 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_97); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_98 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_97); x_99 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__31; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_100 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_99); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_100 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_99); x_101 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__32; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_102 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_101); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_102 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_101); x_103 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__10; x_104 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__11; -lean_inc(x_34); +lean_inc(x_32); +lean_inc(x_29); +x_105 = l_Lean_addMacroScope(x_29, x_104, x_32); lean_inc(x_30); -x_105 = l_Lean_addMacroScope(x_30, x_104, x_34); -lean_inc(x_44); -lean_inc(x_36); +lean_inc(x_45); x_106 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_106, 0, x_36); +lean_ctor_set(x_106, 0, x_45); lean_ctor_set(x_106, 1, x_103); lean_ctor_set(x_106, 2, x_105); -lean_ctor_set(x_106, 3, x_44); -lean_inc(x_28); +lean_ctor_set(x_106, 3, x_30); +lean_inc(x_40); lean_inc(x_102); -lean_inc(x_36); -x_107 = l_Lean_Syntax_node2(x_36, x_102, x_106, x_28); +lean_inc(x_45); +x_107 = l_Lean_Syntax_node2(x_45, x_102, x_106, x_40); x_108 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__36; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_109 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_108); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_109 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_108); x_110 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__9; x_111 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__10; -lean_inc(x_36); +lean_inc(x_45); x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_36); +lean_ctor_set(x_112, 0, x_45); lean_ctor_set(x_112, 1, x_111); -lean_inc(x_36); -x_113 = l_Lean_Syntax_node1(x_36, x_110, x_112); -lean_inc(x_28); +lean_inc(x_45); +x_113 = l_Lean_Syntax_node1(x_45, x_110, x_112); +lean_inc(x_40); lean_inc_ref(x_92); lean_inc(x_109); -lean_inc(x_36); -x_114 = l_Lean_Syntax_node3(x_36, x_109, x_92, x_28, x_113); -lean_inc_n(x_28, 2); -lean_inc(x_39); -lean_inc(x_36); -x_115 = l_Lean_Syntax_node3(x_36, x_39, x_28, x_28, x_114); +lean_inc(x_45); +x_114 = l_Lean_Syntax_node3(x_45, x_109, x_92, x_40, x_113); +lean_inc_n(x_40, 2); +lean_inc(x_42); +lean_inc(x_45); +x_115 = l_Lean_Syntax_node3(x_45, x_42, x_40, x_40, x_114); lean_inc(x_100); -lean_inc(x_36); -x_116 = l_Lean_Syntax_node2(x_36, x_100, x_107, x_115); -lean_inc(x_36); +lean_inc(x_45); +x_116 = l_Lean_Syntax_node2(x_45, x_100, x_107, x_115); +lean_inc(x_45); x_117 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_117, 0, x_36); +lean_ctor_set(x_117, 0, x_45); lean_ctor_set(x_117, 1, x_54); x_118 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__13; x_119 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__14; -x_120 = l_Lean_addMacroScope(x_30, x_119, x_34); -lean_inc(x_36); +x_120 = l_Lean_addMacroScope(x_29, x_119, x_32); +lean_inc(x_45); x_121 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_121, 0, x_36); +lean_ctor_set(x_121, 0, x_45); lean_ctor_set(x_121, 1, x_118); lean_ctor_set(x_121, 2, x_120); -lean_ctor_set(x_121, 3, x_44); -lean_inc(x_28); -lean_inc(x_36); -x_122 = l_Lean_Syntax_node2(x_36, x_102, x_121, x_28); +lean_ctor_set(x_121, 3, x_30); +lean_inc(x_40); +lean_inc(x_45); +x_122 = l_Lean_Syntax_node2(x_45, x_102, x_121, x_40); x_123 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__15; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_124 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_123); -lean_inc(x_36); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_124 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_123); +lean_inc(x_45); x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_36); +lean_ctor_set(x_125, 0, x_45); lean_ctor_set(x_125, 1, x_123); x_126 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__16; -lean_inc_ref(x_41); -lean_inc_ref(x_46); -lean_inc_ref(x_42); -x_127 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_126); -lean_inc(x_39); -lean_inc(x_36); -x_128 = l_Lean_Syntax_node1(x_36, x_39, x_35); +lean_inc_ref(x_31); +lean_inc_ref(x_44); +lean_inc_ref(x_27); +x_127 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_126); +lean_inc(x_42); +lean_inc(x_45); +x_128 = l_Lean_Syntax_node1(x_45, x_42, x_39); x_129 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__17; -lean_inc(x_36); +lean_inc(x_45); x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_36); +lean_ctor_set(x_130, 0, x_45); lean_ctor_set(x_130, 1, x_129); -lean_inc(x_28); -lean_inc(x_36); -x_131 = l_Lean_Syntax_node4(x_36, x_127, x_128, x_28, x_130, x_40); -lean_inc(x_36); -x_132 = l_Lean_Syntax_node2(x_36, x_124, x_125, x_131); -lean_inc(x_28); +lean_inc(x_40); +lean_inc(x_45); +x_131 = l_Lean_Syntax_node4(x_45, x_127, x_128, x_40, x_130, x_46); +lean_inc(x_45); +x_132 = l_Lean_Syntax_node2(x_45, x_124, x_125, x_131); +lean_inc(x_40); lean_inc_ref(x_92); -lean_inc(x_36); -x_133 = l_Lean_Syntax_node3(x_36, x_109, x_92, x_28, x_132); -lean_inc_n(x_28, 2); -lean_inc(x_39); -lean_inc(x_36); -x_134 = l_Lean_Syntax_node3(x_36, x_39, x_28, x_28, x_133); -lean_inc(x_36); -x_135 = l_Lean_Syntax_node2(x_36, x_100, x_122, x_134); -lean_inc(x_39); -lean_inc(x_36); -x_136 = l_Lean_Syntax_node3(x_36, x_39, x_116, x_117, x_135); -lean_inc(x_36); -x_137 = l_Lean_Syntax_node1(x_36, x_98, x_136); +lean_inc(x_45); +x_133 = l_Lean_Syntax_node3(x_45, x_109, x_92, x_40, x_132); +lean_inc_n(x_40, 2); +lean_inc(x_42); +lean_inc(x_45); +x_134 = l_Lean_Syntax_node3(x_45, x_42, x_40, x_40, x_133); +lean_inc(x_45); +x_135 = l_Lean_Syntax_node2(x_45, x_100, x_122, x_134); +lean_inc(x_42); +lean_inc(x_45); +x_136 = l_Lean_Syntax_node3(x_45, x_42, x_116, x_117, x_135); +lean_inc(x_45); +x_137 = l_Lean_Syntax_node1(x_45, x_98, x_136); x_138 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__45; -x_139 = l_Lean_Name_mkStr4(x_42, x_46, x_41, x_138); -lean_inc(x_28); -lean_inc(x_36); -x_140 = l_Lean_Syntax_node1(x_36, x_139, x_28); +x_139 = l_Lean_Name_mkStr4(x_27, x_44, x_31, x_138); +lean_inc(x_40); +lean_inc(x_45); +x_140 = l_Lean_Syntax_node1(x_45, x_139, x_40); x_141 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__46; -lean_inc(x_36); +lean_inc(x_45); x_142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_142, 0, x_36); +lean_ctor_set(x_142, 0, x_45); lean_ctor_set(x_142, 1, x_141); -lean_inc_n(x_28, 2); -lean_inc(x_36); -x_143 = l_Lean_Syntax_node6(x_36, x_94, x_96, x_28, x_137, x_140, x_28, x_142); -lean_inc_n(x_28, 2); -lean_inc(x_36); -x_144 = l_Lean_Syntax_node2(x_36, x_45, x_28, x_28); -if (lean_obj_tag(x_32) == 0) +lean_inc_n(x_40, 2); +lean_inc(x_45); +x_143 = l_Lean_Syntax_node6(x_45, x_94, x_96, x_40, x_137, x_140, x_40, x_142); +lean_inc_n(x_40, 2); +lean_inc(x_45); +x_144 = l_Lean_Syntax_node2(x_45, x_43, x_40, x_40); +if (lean_obj_tag(x_41) == 0) { lean_object* x_145; x_145 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_4 = x_66; -x_5 = x_144; -x_6 = x_28; -x_7 = x_92; -x_8 = x_29; -x_9 = x_73; -x_10 = x_62; -x_11 = x_36; -x_12 = x_38; -x_13 = x_143; -x_14 = x_37; -x_15 = x_39; -x_16 = x_43; -x_17 = x_90; -x_18 = x_64; +x_4 = x_33; +x_5 = x_73; +x_6 = x_35; +x_7 = x_62; +x_8 = x_37; +x_9 = x_38; +x_10 = x_144; +x_11 = x_143; +x_12 = x_90; +x_13 = x_64; +x_14 = x_92; +x_15 = x_40; +x_16 = x_42; +x_17 = x_66; +x_18 = x_45; x_19 = x_145; goto block_26; } else { lean_object* x_146; lean_object* x_147; -x_146 = lean_ctor_get(x_32, 0); +x_146 = lean_ctor_get(x_41, 0); lean_inc(x_146); -lean_dec_ref(x_32); +lean_dec_ref(x_41); x_147 = l_Array_mkArray1___redArg(x_146); -x_4 = x_66; -x_5 = x_144; -x_6 = x_28; -x_7 = x_92; -x_8 = x_29; -x_9 = x_73; -x_10 = x_62; -x_11 = x_36; -x_12 = x_38; -x_13 = x_143; -x_14 = x_37; -x_15 = x_39; -x_16 = x_43; -x_17 = x_90; -x_18 = x_64; +x_4 = x_33; +x_5 = x_73; +x_6 = x_35; +x_7 = x_62; +x_8 = x_37; +x_9 = x_38; +x_10 = x_144; +x_11 = x_143; +x_12 = x_90; +x_13 = x_64; +x_14 = x_92; +x_15 = x_40; +x_16 = x_42; +x_17 = x_66; +x_18 = x_45; x_19 = x_147; goto block_26; } @@ -5747,7 +5747,7 @@ lean_inc(x_168); lean_inc(x_166); lean_inc(x_165); lean_ctor_set(x_162, 5, x_168); -x_169 = l_Lake_DSL_expandOptSimpleBinder(x_150, x_162, x_163); +x_169 = l_Lake_DSL_expandOptSimpleBinder(x_154, x_162, x_163); if (lean_obj_tag(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; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; @@ -5760,13 +5760,13 @@ x_172 = l_Lean_SourceInfo_fromRef(x_168, x_158); lean_dec(x_168); x_173 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__20; x_174 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__47; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_175 = l_Lean_Name_mkStr4(x_155, x_160, x_173, x_174); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_175 = l_Lean_Name_mkStr4(x_149, x_157, x_173, x_174); x_176 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__49; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_177 = l_Lean_Name_mkStr4(x_155, x_160, x_173, x_176); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_177 = l_Lean_Name_mkStr4(x_149, x_157, x_173, x_176); x_178 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__9; x_179 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__10; lean_inc(x_172); @@ -5779,9 +5779,9 @@ lean_inc(x_172); x_181 = l_Lean_Syntax_node1(x_172, x_177, x_180); x_182 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__51; x_183 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__52; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_184 = l_Lean_Name_mkStr4(x_155, x_160, x_182, x_183); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_184 = l_Lean_Name_mkStr4(x_149, x_157, x_182, x_183); x_185 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__19; x_186 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__21; lean_inc(x_166); @@ -5799,44 +5799,44 @@ lean_inc(x_172); x_190 = l_Lean_Syntax_node2(x_172, x_184, x_189, x_180); lean_inc(x_172); x_191 = l_Lean_Syntax_node2(x_172, x_175, x_181, x_190); -x_192 = lean_mk_empty_array_with_capacity(x_154); +x_192 = lean_mk_empty_array_with_capacity(x_159); x_193 = lean_array_push(x_192, x_191); -x_194 = l_Lake_DSL_expandAttrs(x_149); +x_194 = l_Lake_DSL_expandAttrs(x_155); x_195 = l_Array_append___redArg(x_193, x_194); lean_dec_ref(x_194); x_196 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__4; -lean_inc_ref(x_156); -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_197 = l_Lean_Name_mkStr4(x_155, x_160, x_156, x_196); +lean_inc_ref(x_151); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_197 = l_Lean_Name_mkStr4(x_149, x_157, x_151, x_196); x_198 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__6; -lean_inc_ref(x_156); -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_199 = l_Lean_Name_mkStr4(x_155, x_160, x_156, x_198); +lean_inc_ref(x_151); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_199 = l_Lean_Name_mkStr4(x_149, x_157, x_151, x_198); if (lean_obj_tag(x_153) == 0) { lean_object* x_200; x_200 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_27 = x_199; -x_28 = x_180; -x_29 = x_179; -x_30 = x_165; -x_31 = x_195; -x_32 = x_161; -x_33 = x_156; -x_34 = x_166; -x_35 = x_170; -x_36 = x_172; +x_27 = x_149; +x_28 = x_151; +x_29 = x_165; +x_30 = x_188; +x_31 = x_173; +x_32 = x_166; +x_33 = x_179; +x_34 = x_199; +x_35 = x_171; +x_36 = x_195; x_37 = x_197; -x_38 = x_171; -x_39 = x_178; -x_40 = x_151; -x_41 = x_173; -x_42 = x_155; -x_43 = x_157; -x_44 = x_188; -x_45 = x_159; +x_38 = x_150; +x_39 = x_170; +x_40 = x_180; +x_41 = x_161; +x_42 = x_178; +x_43 = x_156; +x_44 = x_157; +x_45 = x_172; x_46 = x_160; x_47 = x_200; goto block_148; @@ -5848,25 +5848,25 @@ x_201 = lean_ctor_get(x_153, 0); lean_inc(x_201); lean_dec_ref(x_153); x_202 = l_Array_mkArray1___redArg(x_201); -x_27 = x_199; -x_28 = x_180; -x_29 = x_179; -x_30 = x_165; -x_31 = x_195; -x_32 = x_161; -x_33 = x_156; -x_34 = x_166; -x_35 = x_170; -x_36 = x_172; +x_27 = x_149; +x_28 = x_151; +x_29 = x_165; +x_30 = x_188; +x_31 = x_173; +x_32 = x_166; +x_33 = x_179; +x_34 = x_199; +x_35 = x_171; +x_36 = x_195; x_37 = x_197; -x_38 = x_171; -x_39 = x_178; -x_40 = x_151; -x_41 = x_173; -x_42 = x_155; -x_43 = x_157; -x_44 = x_188; -x_45 = x_159; +x_38 = x_150; +x_39 = x_170; +x_40 = x_180; +x_41 = x_161; +x_42 = x_178; +x_43 = x_156; +x_44 = x_157; +x_45 = x_172; x_46 = x_160; x_47 = x_202; goto block_148; @@ -5879,14 +5879,14 @@ lean_dec(x_168); lean_dec(x_166); lean_dec(x_165); lean_dec(x_161); -lean_dec_ref(x_160); -lean_dec(x_159); -lean_dec(x_157); -lean_dec_ref(x_156); -lean_dec_ref(x_155); +lean_dec(x_160); +lean_dec_ref(x_157); +lean_dec(x_156); +lean_dec(x_155); lean_dec(x_153); -lean_dec(x_151); -lean_dec(x_149); +lean_dec_ref(x_151); +lean_dec(x_150); +lean_dec_ref(x_149); x_203 = !lean_is_exclusive(x_169); if (x_203 == 0) { @@ -5936,7 +5936,7 @@ lean_ctor_set(x_214, 2, x_209); lean_ctor_set(x_214, 3, x_210); lean_ctor_set(x_214, 4, x_211); lean_ctor_set(x_214, 5, x_213); -x_215 = l_Lake_DSL_expandOptSimpleBinder(x_150, x_214, x_163); +x_215 = l_Lake_DSL_expandOptSimpleBinder(x_154, x_214, x_163); if (lean_obj_tag(x_215) == 0) { lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; @@ -5949,13 +5949,13 @@ x_218 = l_Lean_SourceInfo_fromRef(x_213, x_158); lean_dec(x_213); x_219 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__20; x_220 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__47; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_221 = l_Lean_Name_mkStr4(x_155, x_160, x_219, x_220); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_221 = l_Lean_Name_mkStr4(x_149, x_157, x_219, x_220); x_222 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__49; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_223 = l_Lean_Name_mkStr4(x_155, x_160, x_219, x_222); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_223 = l_Lean_Name_mkStr4(x_149, x_157, x_219, x_222); x_224 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__9; x_225 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__10; lean_inc(x_218); @@ -5968,9 +5968,9 @@ lean_inc(x_218); x_227 = l_Lean_Syntax_node1(x_218, x_223, x_226); x_228 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__51; x_229 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__52; -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_230 = l_Lean_Name_mkStr4(x_155, x_160, x_228, x_229); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_230 = l_Lean_Name_mkStr4(x_149, x_157, x_228, x_229); x_231 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__19; x_232 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__21; lean_inc(x_209); @@ -5988,44 +5988,44 @@ lean_inc(x_218); x_236 = l_Lean_Syntax_node2(x_218, x_230, x_235, x_226); lean_inc(x_218); x_237 = l_Lean_Syntax_node2(x_218, x_221, x_227, x_236); -x_238 = lean_mk_empty_array_with_capacity(x_154); +x_238 = lean_mk_empty_array_with_capacity(x_159); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lake_DSL_expandAttrs(x_149); +x_240 = l_Lake_DSL_expandAttrs(x_155); x_241 = l_Array_append___redArg(x_239, x_240); lean_dec_ref(x_240); x_242 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__4; -lean_inc_ref(x_156); -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_243 = l_Lean_Name_mkStr4(x_155, x_160, x_156, x_242); +lean_inc_ref(x_151); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_243 = l_Lean_Name_mkStr4(x_149, x_157, x_151, x_242); x_244 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__6; -lean_inc_ref(x_156); -lean_inc_ref(x_160); -lean_inc_ref(x_155); -x_245 = l_Lean_Name_mkStr4(x_155, x_160, x_156, x_244); +lean_inc_ref(x_151); +lean_inc_ref(x_157); +lean_inc_ref(x_149); +x_245 = l_Lean_Name_mkStr4(x_149, x_157, x_151, x_244); if (lean_obj_tag(x_153) == 0) { lean_object* x_246; x_246 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_27 = x_245; -x_28 = x_226; -x_29 = x_225; -x_30 = x_208; -x_31 = x_241; -x_32 = x_161; -x_33 = x_156; -x_34 = x_209; -x_35 = x_216; -x_36 = x_218; +x_27 = x_149; +x_28 = x_151; +x_29 = x_208; +x_30 = x_234; +x_31 = x_219; +x_32 = x_209; +x_33 = x_225; +x_34 = x_245; +x_35 = x_217; +x_36 = x_241; x_37 = x_243; -x_38 = x_217; -x_39 = x_224; -x_40 = x_151; -x_41 = x_219; -x_42 = x_155; -x_43 = x_157; -x_44 = x_234; -x_45 = x_159; +x_38 = x_150; +x_39 = x_216; +x_40 = x_226; +x_41 = x_161; +x_42 = x_224; +x_43 = x_156; +x_44 = x_157; +x_45 = x_218; x_46 = x_160; x_47 = x_246; goto block_148; @@ -6037,25 +6037,25 @@ x_247 = lean_ctor_get(x_153, 0); lean_inc(x_247); lean_dec_ref(x_153); x_248 = l_Array_mkArray1___redArg(x_247); -x_27 = x_245; -x_28 = x_226; -x_29 = x_225; -x_30 = x_208; -x_31 = x_241; -x_32 = x_161; -x_33 = x_156; -x_34 = x_209; -x_35 = x_216; -x_36 = x_218; +x_27 = x_149; +x_28 = x_151; +x_29 = x_208; +x_30 = x_234; +x_31 = x_219; +x_32 = x_209; +x_33 = x_225; +x_34 = x_245; +x_35 = x_217; +x_36 = x_241; x_37 = x_243; -x_38 = x_217; -x_39 = x_224; -x_40 = x_151; -x_41 = x_219; -x_42 = x_155; -x_43 = x_157; -x_44 = x_234; -x_45 = x_159; +x_38 = x_150; +x_39 = x_216; +x_40 = x_226; +x_41 = x_161; +x_42 = x_224; +x_43 = x_156; +x_44 = x_157; +x_45 = x_218; x_46 = x_160; x_47 = x_248; goto block_148; @@ -6068,14 +6068,14 @@ lean_dec(x_213); lean_dec(x_209); lean_dec(x_208); lean_dec(x_161); -lean_dec_ref(x_160); -lean_dec(x_159); -lean_dec(x_157); -lean_dec_ref(x_156); -lean_dec_ref(x_155); +lean_dec(x_160); +lean_dec_ref(x_157); +lean_dec(x_156); +lean_dec(x_155); lean_dec(x_153); -lean_dec(x_151); -lean_dec(x_149); +lean_dec_ref(x_151); +lean_dec(x_150); +lean_dec_ref(x_149); x_249 = lean_ctor_get(x_215, 0); lean_inc(x_249); x_250 = lean_ctor_get(x_215, 1); @@ -6102,100 +6102,100 @@ return x_252; block_273: { lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_268 = l_Array_append___redArg(x_260, x_267); +x_268 = l_Array_append___redArg(x_255, x_267); lean_dec_ref(x_267); -lean_inc(x_264); +lean_inc(x_262); x_269 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_269, 0, x_264); -lean_ctor_set(x_269, 1, x_263); +lean_ctor_set(x_269, 0, x_262); +lean_ctor_set(x_269, 1, x_258); lean_ctor_set(x_269, 2, x_268); -lean_inc(x_264); -x_270 = l_Lean_Syntax_node4(x_264, x_255, x_259, x_266, x_265, x_269); -x_271 = l_Lean_Syntax_node5(x_264, x_254, x_261, x_262, x_258, x_257, x_270); +lean_inc(x_262); +x_270 = l_Lean_Syntax_node4(x_262, x_261, x_256, x_260, x_263, x_269); +x_271 = l_Lean_Syntax_node5(x_262, x_254, x_265, x_257, x_266, x_259, x_270); x_272 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_256); +lean_ctor_set(x_272, 1, x_264); return x_272; } block_305: { lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; -lean_inc_ref(x_274); -x_288 = l_Array_append___redArg(x_274, x_287); +lean_inc_ref(x_279); +x_288 = l_Array_append___redArg(x_279, x_287); lean_dec_ref(x_287); -lean_inc(x_286); -lean_inc(x_277); +lean_inc(x_283); +lean_inc(x_276); x_289 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_289, 0, x_277); -lean_ctor_set(x_289, 1, x_286); +lean_ctor_set(x_289, 0, x_276); +lean_ctor_set(x_289, 1, x_283); lean_ctor_set(x_289, 2, x_288); x_290 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand_spec__6___lam__0___closed__3; x_291 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__25; -lean_inc_ref(x_275); -lean_inc_ref(x_282); -x_292 = l_Lean_Name_mkStr4(x_282, x_275, x_290, x_291); +lean_inc_ref(x_278); +lean_inc_ref(x_284); +x_292 = l_Lean_Name_mkStr4(x_284, x_278, x_290, x_291); x_293 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__37; -lean_inc(x_277); +lean_inc(x_276); x_294 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_294, 0, x_277); +lean_ctor_set(x_294, 0, x_276); lean_ctor_set(x_294, 1, x_293); -lean_inc(x_277); +lean_inc(x_276); x_295 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_295, 0, x_277); -lean_ctor_set(x_295, 1, x_279); -lean_inc(x_277); -x_296 = l_Lean_Syntax_node2(x_277, x_283, x_295, x_280); +lean_ctor_set(x_295, 0, x_276); +lean_ctor_set(x_295, 1, x_280); +lean_inc(x_276); +x_296 = l_Lean_Syntax_node2(x_276, x_275, x_295, x_281); x_297 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__26; x_298 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__27; -x_299 = l_Lean_Name_mkStr4(x_282, x_275, x_297, x_298); -lean_inc_ref(x_274); -lean_inc(x_286); -lean_inc(x_277); +x_299 = l_Lean_Name_mkStr4(x_284, x_278, x_297, x_298); +lean_inc_ref(x_279); +lean_inc(x_283); +lean_inc(x_276); x_300 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_300, 0, x_277); -lean_ctor_set(x_300, 1, x_286); -lean_ctor_set(x_300, 2, x_274); +lean_ctor_set(x_300, 0, x_276); +lean_ctor_set(x_300, 1, x_283); +lean_ctor_set(x_300, 2, x_279); lean_inc_ref(x_300); -lean_inc(x_277); -x_301 = l_Lean_Syntax_node2(x_277, x_299, x_300, x_300); -if (lean_obj_tag(x_284) == 0) +lean_inc(x_276); +x_301 = l_Lean_Syntax_node2(x_276, x_299, x_300, x_300); +if (lean_obj_tag(x_274) == 0) { lean_object* x_302; x_302 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_255 = x_292; -x_256 = x_278; -x_257 = x_289; -x_258 = x_281; -x_259 = x_294; -x_260 = x_274; -x_261 = x_276; -x_262 = x_285; -x_263 = x_286; -x_264 = x_277; -x_265 = x_301; -x_266 = x_296; +x_255 = x_279; +x_256 = x_294; +x_257 = x_282; +x_258 = x_283; +x_259 = x_289; +x_260 = x_296; +x_261 = x_292; +x_262 = x_276; +x_263 = x_301; +x_264 = x_285; +x_265 = x_277; +x_266 = x_286; x_267 = x_302; goto block_273; } else { lean_object* x_303; lean_object* x_304; -x_303 = lean_ctor_get(x_284, 0); +x_303 = lean_ctor_get(x_274, 0); lean_inc(x_303); -lean_dec_ref(x_284); +lean_dec_ref(x_274); x_304 = l_Array_mkArray1___redArg(x_303); -x_255 = x_292; -x_256 = x_278; -x_257 = x_289; -x_258 = x_281; -x_259 = x_294; -x_260 = x_274; -x_261 = x_276; -x_262 = x_285; -x_263 = x_286; -x_264 = x_277; -x_265 = x_301; -x_266 = x_296; +x_255 = x_279; +x_256 = x_294; +x_257 = x_282; +x_258 = x_283; +x_259 = x_289; +x_260 = x_296; +x_261 = x_292; +x_262 = x_276; +x_263 = x_301; +x_264 = x_285; +x_265 = x_277; +x_266 = x_286; x_267 = x_304; goto block_273; } @@ -6203,22 +6203,22 @@ goto block_273; block_329: { lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; -lean_inc_ref(x_307); -x_321 = l_Array_append___redArg(x_307, x_320); +lean_inc_ref(x_312); +x_321 = l_Array_append___redArg(x_312, x_320); lean_dec_ref(x_320); -lean_inc(x_319); -lean_inc(x_310); +lean_inc(x_315); +lean_inc(x_309); x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_310); -lean_ctor_set(x_322, 1, x_319); +lean_ctor_set(x_322, 0, x_309); +lean_ctor_set(x_322, 1, x_315); lean_ctor_set(x_322, 2, x_321); -x_323 = l_Lean_SourceInfo_fromRef(x_313, x_306); -lean_dec(x_313); +x_323 = l_Lean_SourceInfo_fromRef(x_316, x_306); +lean_dec(x_316); x_324 = l___private_Lake_DSL_Package_0__Lake_DSL_expandPostUpdateDecl___closed__20; x_325 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_325, 0, x_323); lean_ctor_set(x_325, 1, x_324); -if (lean_obj_tag(x_312) == 0) +if (lean_obj_tag(x_317) == 0) { lean_object* x_326; x_326 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; @@ -6227,37 +6227,37 @@ x_275 = x_308; x_276 = x_309; x_277 = x_310; x_278 = x_311; -x_279 = x_314; -x_280 = x_315; -x_281 = x_325; -x_282 = x_316; -x_283 = x_317; +x_279 = x_312; +x_280 = x_313; +x_281 = x_314; +x_282 = x_322; +x_283 = x_315; x_284 = x_318; -x_285 = x_322; -x_286 = x_319; +x_285 = x_319; +x_286 = x_325; x_287 = x_326; goto block_305; } else { lean_object* x_327; lean_object* x_328; -x_327 = lean_ctor_get(x_312, 0); +x_327 = lean_ctor_get(x_317, 0); lean_inc(x_327); -lean_dec_ref(x_312); +lean_dec_ref(x_317); x_328 = l_Array_mkArray1___redArg(x_327); x_274 = x_307; x_275 = x_308; x_276 = x_309; x_277 = x_310; x_278 = x_311; -x_279 = x_314; -x_280 = x_315; -x_281 = x_325; -x_282 = x_316; -x_283 = x_317; +x_279 = x_312; +x_280 = x_313; +x_281 = x_314; +x_282 = x_322; +x_283 = x_315; x_284 = x_318; -x_285 = x_322; -x_286 = x_319; +x_285 = x_319; +x_286 = x_325; x_287 = x_328; goto block_305; } @@ -6265,23 +6265,23 @@ goto block_305; block_349: { lean_object* x_344; lean_object* x_345; -lean_inc_ref(x_331); -x_344 = l_Array_append___redArg(x_331, x_343); +lean_inc_ref(x_335); +x_344 = l_Array_append___redArg(x_335, x_343); lean_dec_ref(x_343); -lean_inc(x_342); -lean_inc(x_333); +lean_inc(x_338); +lean_inc(x_332); x_345 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_345, 0, x_333); -lean_ctor_set(x_345, 1, x_342); +lean_ctor_set(x_345, 0, x_332); +lean_ctor_set(x_345, 1, x_338); lean_ctor_set(x_345, 2, x_344); -if (lean_obj_tag(x_330) == 0) +if (lean_obj_tag(x_333) == 0) { lean_object* x_346; x_346 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_307 = x_331; -x_308 = x_332; -x_309 = x_345; -x_310 = x_333; +x_307 = x_330; +x_308 = x_331; +x_309 = x_332; +x_310 = x_345; x_311 = x_334; x_312 = x_335; x_313 = x_336; @@ -6297,14 +6297,14 @@ goto block_329; else { lean_object* x_347; lean_object* x_348; -x_347 = lean_ctor_get(x_330, 0); +x_347 = lean_ctor_get(x_333, 0); lean_inc(x_347); -lean_dec_ref(x_330); +lean_dec_ref(x_333); x_348 = l_Array_mkArray1___redArg(x_347); -x_307 = x_331; -x_308 = x_332; -x_309 = x_345; -x_310 = x_333; +x_307 = x_330; +x_308 = x_331; +x_309 = x_332; +x_310 = x_345; x_311 = x_334; x_312 = x_335; x_313 = x_336; @@ -6333,19 +6333,19 @@ if (lean_obj_tag(x_355) == 0) { lean_object* x_367; x_367 = l___private_Lake_DSL_Package_0__Lake_DSL_elabPackageCommand___closed__28; -x_330 = x_350; -x_331 = x_366; -x_332 = x_358; -x_333 = x_364; -x_334 = x_361; -x_335 = x_351; -x_336 = x_354; -x_337 = x_353; -x_338 = x_352; -x_339 = x_357; +x_330 = x_359; +x_331 = x_353; +x_332 = x_364; +x_333 = x_358; +x_334 = x_350; +x_335 = x_366; +x_336 = x_351; +x_337 = x_352; +x_338 = x_365; +x_339 = x_354; x_340 = x_356; -x_341 = x_359; -x_342 = x_365; +x_341 = x_357; +x_342 = x_361; x_343 = x_367; goto block_349; } @@ -6356,19 +6356,19 @@ x_368 = lean_ctor_get(x_355, 0); lean_inc(x_368); lean_dec_ref(x_355); x_369 = l_Array_mkArray1___redArg(x_368); -x_330 = x_350; -x_331 = x_366; -x_332 = x_358; -x_333 = x_364; -x_334 = x_361; -x_335 = x_351; -x_336 = x_354; -x_337 = x_353; -x_338 = x_352; -x_339 = x_357; +x_330 = x_359; +x_331 = x_353; +x_332 = x_364; +x_333 = x_358; +x_334 = x_350; +x_335 = x_366; +x_336 = x_351; +x_337 = x_352; +x_338 = x_365; +x_339 = x_354; x_340 = x_356; -x_341 = x_359; -x_342 = x_365; +x_341 = x_357; +x_342 = x_361; x_343 = x_369; goto block_349; } diff --git a/stage0/stdlib/Lake/DSL/Require.c b/stage0/stdlib/Lake/DSL/Require.c index 2c7bfa3bacd6..01fd96f65c47 100644 --- a/stage0/stdlib/Lake/DSL/Require.c +++ b/stage0/stdlib/Lake/DSL/Require.c @@ -824,14 +824,14 @@ goto block_91; { uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; x_19 = 0; -x_20 = l_Lean_SourceInfo_fromRef(x_13, x_19); -lean_dec(x_13); +x_20 = l_Lean_SourceInfo_fromRef(x_15, x_19); +lean_dec(x_15); x_21 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__4; x_22 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__1; x_23 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__2; x_24 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__3; x_25 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__4; -x_26 = l_Lean_addMacroScope(x_16, x_25, x_14); +x_26 = l_Lean_addMacroScope(x_14, x_25, x_16); x_27 = l_Lean_Name_mkStr3(x_1, x_23, x_24); x_28 = lean_box(0); lean_inc(x_27); @@ -854,7 +854,7 @@ lean_ctor_set(x_33, 2, x_26); lean_ctor_set(x_33, 3, x_32); x_34 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__12; lean_inc(x_20); -x_35 = l_Lean_Syntax_node3(x_20, x_34, x_2, x_15, x_17); +x_35 = l_Lean_Syntax_node3(x_20, x_34, x_2, x_13, x_17); x_36 = l_Lean_Syntax_node2(x_20, x_21, x_33, x_35); x_37 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_37, 0, x_36); @@ -863,26 +863,26 @@ return x_37; } block_65: { -if (lean_obj_tag(x_42) == 0) +if (lean_obj_tag(x_39) == 0) { uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; x_45 = 0; -x_46 = l_Lean_SourceInfo_fromRef(x_39, x_45); +x_46 = l_Lean_SourceInfo_fromRef(x_41, x_45); x_47 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__1___closed__1; x_48 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__1___closed__2; +lean_inc(x_42); lean_inc(x_40); -lean_inc(x_41); -x_49 = l_Lean_addMacroScope(x_41, x_48, x_40); +x_49 = l_Lean_addMacroScope(x_40, x_48, x_42); x_50 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__1___closed__6; x_51 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_51, 0, x_46); lean_ctor_set(x_51, 1, x_47); lean_ctor_set(x_51, 2, x_49); lean_ctor_set(x_51, 3, x_50); -x_13 = x_39; +x_13 = x_43; x_14 = x_40; -x_15 = x_43; -x_16 = x_41; +x_15 = x_41; +x_16 = x_42; x_17 = x_51; x_18 = x_44; goto block_38; @@ -890,19 +890,19 @@ goto block_38; 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_ctor_get(x_42, 0); +x_52 = lean_ctor_get(x_39, 0); lean_inc(x_52); -lean_dec_ref(x_42); -x_53 = l_Lean_replaceRef(x_52, x_39); +lean_dec_ref(x_39); +x_53 = l_Lean_replaceRef(x_52, x_41); x_54 = 0; x_55 = l_Lean_SourceInfo_fromRef(x_53, x_54); lean_dec(x_53); x_56 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__4; x_57 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__6; x_58 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__7; +lean_inc(x_42); lean_inc(x_40); -lean_inc(x_41); -x_59 = l_Lean_addMacroScope(x_41, x_58, x_40); +x_59 = l_Lean_addMacroScope(x_40, x_58, x_42); x_60 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5___closed__10; lean_inc(x_55); x_61 = lean_alloc_ctor(3, 4, 0); @@ -914,10 +914,10 @@ x_62 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5__ lean_inc(x_55); x_63 = l_Lean_Syntax_node1(x_55, x_62, x_52); x_64 = l_Lean_Syntax_node2(x_55, x_56, x_61, x_63); -x_13 = x_39; +x_13 = x_43; x_14 = x_40; -x_15 = x_43; -x_16 = x_41; +x_15 = x_41; +x_16 = x_42; x_17 = x_64; x_18 = x_44; goto block_38; @@ -951,10 +951,10 @@ lean_ctor_set(x_77, 0, x_72); lean_ctor_set(x_77, 1, x_73); lean_ctor_set(x_77, 2, x_75); lean_ctor_set(x_77, 3, x_76); -x_39 = x_70; -x_40 = x_68; -x_41 = x_67; -x_42 = x_66; +x_39 = x_66; +x_40 = x_67; +x_41 = x_70; +x_42 = x_68; x_43 = x_77; x_44 = x_12; goto block_65; @@ -986,10 +986,10 @@ x_88 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5__ lean_inc(x_81); x_89 = l_Lean_Syntax_node1(x_81, x_88, x_78); x_90 = l_Lean_Syntax_node2(x_81, x_82, x_87, x_89); -x_39 = x_70; -x_40 = x_68; -x_41 = x_67; -x_42 = x_66; +x_39 = x_66; +x_40 = x_67; +x_41 = x_70; +x_42 = x_68; x_43 = x_90; x_44 = x_12; goto block_65; @@ -2198,98 +2198,98 @@ goto block_533; block_193: { lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_91 = l_Array_append___redArg(x_79, x_90); +x_91 = l_Array_append___redArg(x_87, x_90); lean_dec_ref(x_90); +lean_inc(x_71); lean_inc(x_78); -lean_inc(x_84); x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_84); -lean_ctor_set(x_92, 1, x_78); +lean_ctor_set(x_92, 0, x_78); +lean_ctor_set(x_92, 1, x_71); lean_ctor_set(x_92, 2, x_91); x_93 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__8; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_94 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_93); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_94 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_93); x_95 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__9; -lean_inc(x_84); +lean_inc(x_78); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_84); +lean_ctor_set(x_96, 0, x_78); lean_ctor_set(x_96, 1, x_95); x_97 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__10; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_98 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_97); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_98 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_97); x_99 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__11; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_100 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_99); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_100 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_99); lean_inc(x_81); -lean_inc(x_84); -x_101 = l_Lean_Syntax_node1(x_84, x_100, x_81); +lean_inc(x_78); +x_101 = l_Lean_Syntax_node1(x_78, x_100, x_81); x_102 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__12; x_103 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__13; -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_104 = l_Lean_Name_mkStr4(x_88, x_72, x_102, x_103); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_104 = l_Lean_Name_mkStr4(x_70, x_68, x_102, x_103); x_105 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__15; x_106 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__16; +lean_inc(x_67); lean_inc(x_85); -lean_inc(x_75); -x_107 = l_Lean_addMacroScope(x_75, x_106, x_85); +x_107 = l_Lean_addMacroScope(x_85, x_106, x_67); x_108 = lean_box(0); -lean_inc(x_84); +lean_inc(x_78); x_109 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_109, 0, x_84); +lean_ctor_set(x_109, 0, x_78); lean_ctor_set(x_109, 1, x_105); lean_ctor_set(x_109, 2, x_107); lean_ctor_set(x_109, 3, x_108); lean_inc(x_81); -lean_inc(x_84); -x_110 = l_Lean_Syntax_node2(x_84, x_104, x_109, x_81); -lean_inc(x_84); -x_111 = l_Lean_Syntax_node2(x_84, x_98, x_101, x_110); lean_inc(x_78); -lean_inc(x_84); -x_112 = l_Lean_Syntax_node1(x_84, x_78, x_111); +x_110 = l_Lean_Syntax_node2(x_78, x_104, x_109, x_81); +lean_inc(x_78); +x_111 = l_Lean_Syntax_node2(x_78, x_98, x_101, x_110); +lean_inc(x_71); +lean_inc(x_78); +x_112 = l_Lean_Syntax_node1(x_78, x_71, x_111); x_113 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__17; -lean_inc(x_84); +lean_inc(x_78); x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_84); +lean_ctor_set(x_114, 0, x_78); lean_ctor_set(x_114, 1, x_113); -lean_inc(x_84); -x_115 = l_Lean_Syntax_node3(x_84, x_94, x_96, x_112, x_114); lean_inc(x_78); -lean_inc(x_84); -x_116 = l_Lean_Syntax_node1(x_84, x_78, x_115); +x_115 = l_Lean_Syntax_node3(x_78, x_94, x_96, x_112, x_114); +lean_inc(x_71); +lean_inc(x_78); +x_116 = l_Lean_Syntax_node1(x_78, x_71, x_115); lean_inc_n(x_81, 5); -lean_inc(x_84); -x_117 = l_Lean_Syntax_node7(x_84, x_65, x_92, x_116, x_81, x_81, x_81, x_81, x_81); +lean_inc(x_78); +x_117 = l_Lean_Syntax_node7(x_78, x_66, x_92, x_116, x_81, x_81, x_81, x_81, x_81); x_118 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__18; -lean_inc_ref(x_74); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_119 = l_Lean_Name_mkStr4(x_88, x_72, x_74, x_118); +lean_inc_ref(x_80); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_119 = l_Lean_Name_mkStr4(x_70, x_68, x_80, x_118); x_120 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__19; -lean_inc(x_84); +lean_inc(x_78); x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_84); +lean_ctor_set(x_121, 0, x_78); lean_ctor_set(x_121, 1, x_120); x_122 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__20; -lean_inc_ref(x_74); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_123 = l_Lean_Name_mkStr4(x_88, x_72, x_74, x_122); +lean_inc_ref(x_80); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_123 = l_Lean_Name_mkStr4(x_70, x_68, x_80, x_122); x_124 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__21; x_125 = lean_box(2); -lean_inc(x_78); +lean_inc(x_71); x_126 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_78); +lean_ctor_set(x_126, 1, x_71); lean_ctor_set(x_126, 2, x_124); -x_127 = lean_mk_empty_array_with_capacity(x_68); +x_127 = lean_mk_empty_array_with_capacity(x_83); lean_inc(x_77); x_128 = lean_array_push(x_127, x_77); x_129 = lean_array_push(x_128, x_126); @@ -2298,246 +2298,246 @@ lean_ctor_set(x_130, 0, x_125); lean_ctor_set(x_130, 1, x_123); lean_ctor_set(x_130, 2, x_129); x_131 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__22; -lean_inc_ref(x_74); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_132 = l_Lean_Name_mkStr4(x_88, x_72, x_74, x_131); +lean_inc_ref(x_80); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_132 = l_Lean_Name_mkStr4(x_70, x_68, x_80, x_131); x_133 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__23; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_134 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_133); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_134 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_133); x_135 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__24; -lean_inc(x_84); +lean_inc(x_78); x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_84); +lean_ctor_set(x_136, 0, x_78); lean_ctor_set(x_136, 1, x_135); x_137 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__27; -lean_inc(x_84); -x_138 = l_Lean_Syntax_node2(x_84, x_134, x_136, x_137); lean_inc(x_78); -lean_inc(x_84); -x_139 = l_Lean_Syntax_node1(x_84, x_78, x_138); +x_138 = l_Lean_Syntax_node2(x_78, x_134, x_136, x_137); +lean_inc(x_71); +lean_inc(x_78); +x_139 = l_Lean_Syntax_node1(x_78, x_71, x_138); lean_inc(x_81); -lean_inc(x_84); -x_140 = l_Lean_Syntax_node2(x_84, x_132, x_81, x_139); +lean_inc(x_78); +x_140 = l_Lean_Syntax_node2(x_78, x_132, x_81, x_139); x_141 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__28; -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_142 = l_Lean_Name_mkStr4(x_88, x_72, x_74, x_141); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_142 = l_Lean_Name_mkStr4(x_70, x_68, x_80, x_141); x_143 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__29; -lean_inc(x_84); +lean_inc(x_78); x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_84); +lean_ctor_set(x_144, 0, x_78); lean_ctor_set(x_144, 1, x_143); x_145 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__30; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_146 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_145); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_146 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_145); x_147 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__31; -lean_inc_ref(x_83); -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_148 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_147); +lean_inc_ref(x_86); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_148 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_147); x_149 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__33; x_150 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__34; +lean_inc(x_67); lean_inc(x_85); -lean_inc(x_75); -x_151 = l_Lean_addMacroScope(x_75, x_150, x_85); -lean_inc(x_84); +x_151 = l_Lean_addMacroScope(x_85, x_150, x_67); +lean_inc(x_78); x_152 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_152, 0, x_84); +lean_ctor_set(x_152, 0, x_78); lean_ctor_set(x_152, 1, x_149); lean_ctor_set(x_152, 2, x_151); lean_ctor_set(x_152, 3, x_108); lean_inc(x_81); lean_inc(x_148); -lean_inc(x_84); -x_153 = l_Lean_Syntax_node2(x_84, x_148, x_152, x_81); +lean_inc(x_78); +x_153 = l_Lean_Syntax_node2(x_78, x_148, x_152, x_81); x_154 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__35; -lean_inc_ref(x_72); -lean_inc_ref(x_88); -x_155 = l_Lean_Name_mkStr4(x_88, x_72, x_83, x_154); +lean_inc_ref(x_68); +lean_inc_ref(x_70); +x_155 = l_Lean_Name_mkStr4(x_70, x_68, x_86, x_154); x_156 = l_Lean_TSyntax_getId(x_77); lean_dec(x_77); x_157 = l_Lean_instQuoteNameMkStr1___private__1(x_156); lean_inc(x_81); lean_inc_ref(x_144); lean_inc(x_155); -lean_inc(x_84); -x_158 = l_Lean_Syntax_node3(x_84, x_155, x_144, x_81, x_157); +lean_inc(x_78); +x_158 = l_Lean_Syntax_node3(x_78, x_155, x_144, x_81, x_157); lean_inc_n(x_81, 2); +lean_inc(x_71); lean_inc(x_78); -lean_inc(x_84); -x_159 = l_Lean_Syntax_node3(x_84, x_78, x_81, x_81, x_158); +x_159 = l_Lean_Syntax_node3(x_78, x_71, x_81, x_81, x_158); lean_inc(x_146); -lean_inc(x_84); -x_160 = l_Lean_Syntax_node2(x_84, x_146, x_153, x_159); +lean_inc(x_78); +x_160 = l_Lean_Syntax_node2(x_78, x_146, x_153, x_159); x_161 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__36; -lean_inc(x_84); +lean_inc(x_78); x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_84); +lean_ctor_set(x_162, 0, x_78); lean_ctor_set(x_162, 1, x_161); x_163 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__38; x_164 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__39; +lean_inc(x_67); lean_inc(x_85); -lean_inc(x_75); -x_165 = l_Lean_addMacroScope(x_75, x_164, x_85); -lean_inc(x_84); +x_165 = l_Lean_addMacroScope(x_85, x_164, x_67); +lean_inc(x_78); x_166 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_166, 0, x_84); +lean_ctor_set(x_166, 0, x_78); lean_ctor_set(x_166, 1, x_163); lean_ctor_set(x_166, 2, x_165); lean_ctor_set(x_166, 3, x_108); lean_inc(x_81); lean_inc(x_148); -lean_inc(x_84); -x_167 = l_Lean_Syntax_node2(x_84, x_148, x_166, x_81); +lean_inc(x_78); +x_167 = l_Lean_Syntax_node2(x_78, x_148, x_166, x_81); lean_inc(x_81); lean_inc_ref(x_144); lean_inc(x_155); -lean_inc(x_84); -x_168 = l_Lean_Syntax_node3(x_84, x_155, x_144, x_81, x_80); +lean_inc(x_78); +x_168 = l_Lean_Syntax_node3(x_78, x_155, x_144, x_81, x_75); lean_inc_n(x_81, 2); +lean_inc(x_71); lean_inc(x_78); -lean_inc(x_84); -x_169 = l_Lean_Syntax_node3(x_84, x_78, x_81, x_81, x_168); +x_169 = l_Lean_Syntax_node3(x_78, x_71, x_81, x_81, x_168); lean_inc(x_146); -lean_inc(x_84); -x_170 = l_Lean_Syntax_node2(x_84, x_146, x_167, x_169); +lean_inc(x_78); +x_170 = l_Lean_Syntax_node2(x_78, x_146, x_167, x_169); x_171 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__41; x_172 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__42; +lean_inc(x_67); lean_inc(x_85); -lean_inc(x_75); -x_173 = l_Lean_addMacroScope(x_75, x_172, x_85); -lean_inc(x_84); +x_173 = l_Lean_addMacroScope(x_85, x_172, x_67); +lean_inc(x_78); x_174 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_174, 0, x_84); +lean_ctor_set(x_174, 0, x_78); lean_ctor_set(x_174, 1, x_171); lean_ctor_set(x_174, 2, x_173); lean_ctor_set(x_174, 3, x_108); lean_inc(x_81); lean_inc(x_148); -lean_inc(x_84); -x_175 = l_Lean_Syntax_node2(x_84, x_148, x_174, x_81); +lean_inc(x_78); +x_175 = l_Lean_Syntax_node2(x_78, x_148, x_174, x_81); lean_inc(x_81); lean_inc_ref(x_144); lean_inc(x_155); -lean_inc(x_84); -x_176 = l_Lean_Syntax_node3(x_84, x_155, x_144, x_81, x_67); +lean_inc(x_78); +x_176 = l_Lean_Syntax_node3(x_78, x_155, x_144, x_81, x_76); lean_inc_n(x_81, 2); +lean_inc(x_71); lean_inc(x_78); -lean_inc(x_84); -x_177 = l_Lean_Syntax_node3(x_84, x_78, x_81, x_81, x_176); +x_177 = l_Lean_Syntax_node3(x_78, x_71, x_81, x_81, x_176); lean_inc(x_146); -lean_inc(x_84); -x_178 = l_Lean_Syntax_node2(x_84, x_146, x_175, x_177); +lean_inc(x_78); +x_178 = l_Lean_Syntax_node2(x_78, x_146, x_175, x_177); x_179 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__44; x_180 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__45; +lean_inc(x_67); lean_inc(x_85); -lean_inc(x_75); -x_181 = l_Lean_addMacroScope(x_75, x_180, x_85); -lean_inc(x_84); +x_181 = l_Lean_addMacroScope(x_85, x_180, x_67); +lean_inc(x_78); x_182 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_182, 0, x_84); +lean_ctor_set(x_182, 0, x_78); lean_ctor_set(x_182, 1, x_179); lean_ctor_set(x_182, 2, x_181); lean_ctor_set(x_182, 3, x_108); lean_inc(x_81); lean_inc(x_148); -lean_inc(x_84); -x_183 = l_Lean_Syntax_node2(x_84, x_148, x_182, x_81); +lean_inc(x_78); +x_183 = l_Lean_Syntax_node2(x_78, x_148, x_182, x_81); lean_inc(x_81); lean_inc_ref(x_144); lean_inc(x_155); -lean_inc(x_84); -x_184 = l_Lean_Syntax_node3(x_84, x_155, x_144, x_81, x_89); +lean_inc(x_78); +x_184 = l_Lean_Syntax_node3(x_78, x_155, x_144, x_81, x_88); lean_inc_n(x_81, 2); +lean_inc(x_71); lean_inc(x_78); -lean_inc(x_84); -x_185 = l_Lean_Syntax_node3(x_84, x_78, x_81, x_81, x_184); +x_185 = l_Lean_Syntax_node3(x_78, x_71, x_81, x_81, x_184); lean_inc(x_146); -lean_inc(x_84); -x_186 = l_Lean_Syntax_node2(x_84, x_146, x_183, x_185); +lean_inc(x_78); +x_186 = l_Lean_Syntax_node2(x_78, x_146, x_183, x_185); x_187 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__47; x_188 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__48; -x_189 = l_Lean_addMacroScope(x_75, x_188, x_85); -lean_inc(x_84); +x_189 = l_Lean_addMacroScope(x_85, x_188, x_67); +lean_inc(x_78); x_190 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_190, 0, x_84); +lean_ctor_set(x_190, 0, x_78); lean_ctor_set(x_190, 1, x_187); lean_ctor_set(x_190, 2, x_189); lean_ctor_set(x_190, 3, x_108); lean_inc(x_81); -lean_inc(x_84); -x_191 = l_Lean_Syntax_node2(x_84, x_148, x_190, x_81); -if (lean_obj_tag(x_70) == 0) -{ -x_5 = x_155; -x_6 = x_69; -x_7 = x_142; -x_8 = x_178; -x_9 = x_71; -x_10 = x_72; -x_11 = x_191; -x_12 = x_73; -x_13 = x_76; -x_14 = x_146; -x_15 = x_160; -x_16 = x_78; -x_17 = x_119; -x_18 = x_170; -x_19 = x_81; -x_20 = x_82; -x_21 = x_130; -x_22 = x_144; -x_23 = x_162; -x_24 = x_84; -x_25 = x_186; -x_26 = x_117; -x_27 = x_121; -x_28 = x_86; -x_29 = x_87; -x_30 = x_140; -x_31 = x_88; -x_32 = x_66; +lean_inc(x_78); +x_191 = l_Lean_Syntax_node2(x_78, x_148, x_190, x_81); +if (lean_obj_tag(x_79) == 0) +{ +x_5 = x_117; +x_6 = x_65; +x_7 = x_68; +x_8 = x_69; +x_9 = x_146; +x_10 = x_70; +x_11 = x_71; +x_12 = x_72; +x_13 = x_162; +x_14 = x_121; +x_15 = x_74; +x_16 = x_140; +x_17 = x_130; +x_18 = x_144; +x_19 = x_78; +x_20 = x_191; +x_21 = x_186; +x_22 = x_81; +x_23 = x_82; +x_24 = x_160; +x_25 = x_178; +x_26 = x_84; +x_27 = x_142; +x_28 = x_155; +x_29 = x_119; +x_30 = x_89; +x_31 = x_170; +x_32 = x_73; goto block_58; } else { lean_object* x_192; -lean_dec(x_66); -x_192 = lean_ctor_get(x_70, 0); +lean_dec(x_73); +x_192 = lean_ctor_get(x_79, 0); lean_inc(x_192); -lean_dec_ref(x_70); -x_5 = x_155; -x_6 = x_69; -x_7 = x_142; -x_8 = x_178; -x_9 = x_71; -x_10 = x_72; -x_11 = x_191; -x_12 = x_73; -x_13 = x_76; -x_14 = x_146; -x_15 = x_160; -x_16 = x_78; -x_17 = x_119; -x_18 = x_170; -x_19 = x_81; -x_20 = x_82; -x_21 = x_130; -x_22 = x_144; -x_23 = x_162; -x_24 = x_84; -x_25 = x_186; -x_26 = x_117; -x_27 = x_121; -x_28 = x_86; -x_29 = x_87; -x_30 = x_140; -x_31 = x_88; +lean_dec_ref(x_79); +x_5 = x_117; +x_6 = x_65; +x_7 = x_68; +x_8 = x_69; +x_9 = x_146; +x_10 = x_70; +x_11 = x_71; +x_12 = x_72; +x_13 = x_162; +x_14 = x_121; +x_15 = x_74; +x_16 = x_140; +x_17 = x_130; +x_18 = x_144; +x_19 = x_78; +x_20 = x_191; +x_21 = x_186; +x_22 = x_81; +x_23 = x_82; +x_24 = x_160; +x_25 = x_178; +x_26 = x_84; +x_27 = x_142; +x_28 = x_155; +x_29 = x_119; +x_30 = x_89; +x_31 = x_170; x_32 = x_192; goto block_58; } @@ -2546,8 +2546,8 @@ goto block_58; { 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; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; x_204 = 0; -x_205 = l_Lean_SourceInfo_fromRef(x_198, x_204); -lean_dec(x_198); +x_205 = l_Lean_SourceInfo_fromRef(x_196, x_204); +lean_dec(x_196); x_206 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__50; x_207 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__52; x_208 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__53; @@ -2598,31 +2598,31 @@ if (lean_obj_tag(x_2) == 0) { lean_object* x_229; x_229 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__67; -x_65 = x_228; -x_66 = x_225; +x_65 = x_216; +x_66 = x_228; x_67 = x_195; -x_68 = x_199; -x_69 = x_211; -x_70 = x_200; -x_71 = x_216; -x_72 = x_214; -x_73 = x_227; -x_74 = x_226; -x_75 = x_196; -x_76 = x_209; -x_77 = x_194; -x_78 = x_217; -x_79 = x_218; -x_80 = x_201; +x_68 = x_214; +x_69 = x_220; +x_70 = x_213; +x_71 = x_217; +x_72 = x_227; +x_73 = x_225; +x_74 = x_211; +x_75 = x_198; +x_76 = x_199; +x_77 = x_200; +x_78 = x_205; +x_79 = x_201; +x_80 = x_226; x_81 = x_219; -x_82 = x_203; -x_83 = x_215; -x_84 = x_205; -x_85 = x_197; -x_86 = x_223; -x_87 = x_220; -x_88 = x_213; -x_89 = x_202; +x_82 = x_223; +x_83 = x_197; +x_84 = x_209; +x_85 = x_194; +x_86 = x_215; +x_87 = x_218; +x_88 = x_202; +x_89 = x_203; x_90 = x_229; goto block_193; } @@ -2633,31 +2633,31 @@ x_230 = lean_ctor_get(x_2, 0); lean_inc(x_230); lean_dec_ref(x_2); x_231 = l_Array_mkArray1___redArg(x_230); -x_65 = x_228; -x_66 = x_225; +x_65 = x_216; +x_66 = x_228; x_67 = x_195; -x_68 = x_199; -x_69 = x_211; -x_70 = x_200; -x_71 = x_216; -x_72 = x_214; -x_73 = x_227; -x_74 = x_226; -x_75 = x_196; -x_76 = x_209; -x_77 = x_194; -x_78 = x_217; -x_79 = x_218; -x_80 = x_201; +x_68 = x_214; +x_69 = x_220; +x_70 = x_213; +x_71 = x_217; +x_72 = x_227; +x_73 = x_225; +x_74 = x_211; +x_75 = x_198; +x_76 = x_199; +x_77 = x_200; +x_78 = x_205; +x_79 = x_201; +x_80 = x_226; x_81 = x_219; -x_82 = x_203; -x_83 = x_215; -x_84 = x_205; -x_85 = x_197; -x_86 = x_223; -x_87 = x_220; -x_88 = x_213; -x_89 = x_202; +x_82 = x_223; +x_83 = x_197; +x_84 = x_209; +x_85 = x_194; +x_86 = x_215; +x_87 = x_218; +x_88 = x_202; +x_89 = x_203; x_90 = x_231; goto block_193; } @@ -2666,7 +2666,7 @@ goto block_193; { lean_object* x_241; x_241 = l_Lake_DSL_expandIdentOrStrAsIdent(x_236); -if (lean_obj_tag(x_233) == 0) +if (lean_obj_tag(x_234) == 0) { lean_object* x_242; lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; x_242 = lean_ctor_get(x_239, 1); @@ -2689,13 +2689,13 @@ lean_ctor_set(x_251, 0, x_246); lean_ctor_set(x_251, 1, x_247); lean_ctor_set(x_251, 2, x_249); lean_ctor_set(x_251, 3, x_250); -x_194 = x_241; -x_195 = x_238; -x_196 = x_242; -x_197 = x_243; -x_198 = x_244; -x_199 = x_234; -x_200 = x_235; +x_194 = x_242; +x_195 = x_243; +x_196 = x_244; +x_197 = x_233; +x_198 = x_235; +x_199 = x_238; +x_200 = x_241; x_201 = x_237; x_202 = x_251; x_203 = x_240; @@ -2704,9 +2704,9 @@ goto block_232; else { lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_252 = lean_ctor_get(x_233, 0); +x_252 = lean_ctor_get(x_234, 0); lean_inc(x_252); -lean_dec_ref(x_233); +lean_dec_ref(x_234); x_253 = lean_ctor_get(x_239, 1); lean_inc(x_253); x_254 = lean_ctor_get(x_239, 2); @@ -2735,13 +2735,13 @@ x_265 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5_ lean_inc(x_258); x_266 = l_Lean_Syntax_node1(x_258, x_265, x_252); x_267 = l_Lean_Syntax_node2(x_258, x_259, x_264, x_266); -x_194 = x_241; -x_195 = x_238; -x_196 = x_253; -x_197 = x_254; -x_198 = x_255; -x_199 = x_234; -x_200 = x_235; +x_194 = x_253; +x_195 = x_254; +x_196 = x_255; +x_197 = x_233; +x_198 = x_235; +x_199 = x_238; +x_200 = x_241; x_201 = x_237; x_202 = x_267; x_203 = x_240; @@ -2758,21 +2758,21 @@ lean_inc(x_276); x_277 = lean_ctor_get(x_275, 1); lean_inc(x_277); lean_dec_ref(x_275); -x_233 = x_270; -x_234 = x_269; -x_235 = x_272; -x_236 = x_271; -x_237 = x_273; +x_233 = x_269; +x_234 = x_271; +x_235 = x_270; +x_236 = x_273; +x_237 = x_274; x_238 = x_276; -x_239 = x_274; +x_239 = x_272; x_240 = x_277; goto block_268; } else { -lean_dec_ref(x_274); +lean_dec(x_274); lean_dec(x_273); -lean_dec(x_272); +lean_dec_ref(x_272); lean_dec(x_271); lean_dec(x_270); lean_dec(x_2); @@ -2781,12 +2781,12 @@ return x_275; } block_363: { -if (lean_obj_tag(x_285) == 0) +if (lean_obj_tag(x_284) == 0) { lean_object* x_289; lean_object* x_290; lean_object* x_291; uint8_t x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_289 = lean_ctor_get(x_287, 1); -x_290 = lean_ctor_get(x_287, 2); -x_291 = lean_ctor_get(x_287, 5); +x_289 = lean_ctor_get(x_282, 1); +x_290 = lean_ctor_get(x_282, 2); +x_291 = lean_ctor_get(x_282, 5); x_292 = 0; x_293 = l_Lean_SourceInfo_fromRef(x_291, x_292); x_294 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__1___closed__1; @@ -2800,28 +2800,28 @@ lean_ctor_set(x_298, 0, x_293); lean_ctor_set(x_298, 1, x_294); lean_ctor_set(x_298, 2, x_296); lean_ctor_set(x_298, 3, x_297); -x_233 = x_282; -x_234 = x_281; -x_235 = x_284; -x_236 = x_283; -x_237 = x_288; +x_233 = x_281; +x_234 = x_283; +x_235 = x_288; +x_236 = x_285; +x_237 = x_287; x_238 = x_298; -x_239 = x_287; +x_239 = x_282; x_240 = x_286; goto block_268; } 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; uint8_t x_307; lean_object* x_308; lean_object* x_309; -x_299 = lean_ctor_get(x_285, 0); +x_299 = lean_ctor_get(x_284, 0); lean_inc(x_299); -lean_dec_ref(x_285); -x_300 = lean_ctor_get(x_287, 0); -x_301 = lean_ctor_get(x_287, 1); -x_302 = lean_ctor_get(x_287, 2); -x_303 = lean_ctor_get(x_287, 3); -x_304 = lean_ctor_get(x_287, 4); -x_305 = lean_ctor_get(x_287, 5); +lean_dec_ref(x_284); +x_300 = lean_ctor_get(x_282, 0); +x_301 = lean_ctor_get(x_282, 1); +x_302 = lean_ctor_get(x_282, 2); +x_303 = lean_ctor_get(x_282, 3); +x_304 = lean_ctor_get(x_282, 4); +x_305 = lean_ctor_get(x_282, 5); x_306 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__69; lean_inc(x_299); x_307 = l_Lean_Syntax_isOfKind(x_299, x_306); @@ -2847,10 +2847,10 @@ x_310 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__70; x_311 = l_Lean_Macro_throwErrorAt___redArg(x_299, x_310, x_309, x_286); lean_dec(x_299); x_269 = x_281; -x_270 = x_282; +x_270 = x_288; x_271 = x_283; -x_272 = x_284; -x_273 = x_288; +x_272 = x_282; +x_273 = x_285; x_274 = x_287; x_275 = x_311; goto block_278; @@ -2873,10 +2873,10 @@ x_315 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__70; x_316 = l_Lean_Macro_throwErrorAt___redArg(x_299, x_315, x_309, x_286); lean_dec(x_299); x_269 = x_281; -x_270 = x_282; +x_270 = x_288; x_271 = x_283; -x_272 = x_284; -x_273 = x_288; +x_272 = x_282; +x_273 = x_285; x_274 = x_287; x_275 = x_316; goto block_278; @@ -2906,13 +2906,13 @@ x_325 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5_ lean_inc(x_318); x_326 = l_Lean_Syntax_node1(x_318, x_325, x_317); x_327 = l_Lean_Syntax_node2(x_318, x_319, x_324, x_326); -x_233 = x_282; -x_234 = x_281; -x_235 = x_284; -x_236 = x_283; -x_237 = x_288; +x_233 = x_281; +x_234 = x_283; +x_235 = x_288; +x_236 = x_285; +x_237 = x_287; x_238 = x_327; -x_239 = x_287; +x_239 = x_282; x_240 = x_286; goto block_268; } @@ -2991,13 +2991,13 @@ x_360 = l_Lean_Syntax_node3(x_330, x_338, x_349, x_357, x_359); lean_inc(x_330); x_361 = l_Lean_Syntax_node1(x_330, x_337, x_360); x_362 = l_Lean_Syntax_node2(x_330, x_331, x_336, x_361); -x_233 = x_282; -x_234 = x_281; -x_235 = x_284; -x_236 = x_283; -x_237 = x_288; +x_233 = x_281; +x_234 = x_283; +x_235 = x_288; +x_236 = x_285; +x_237 = x_287; x_238 = x_362; -x_239 = x_287; +x_239 = x_282; x_240 = x_286; goto block_268; } @@ -3008,17 +3008,17 @@ goto block_268; { uint8_t x_371; lean_inc(x_279); -x_371 = l_Lean_Syntax_isOfKind(x_279, x_364); -lean_dec(x_364); +x_371 = l_Lean_Syntax_isOfKind(x_279, x_368); +lean_dec(x_368); if (x_371 == 0) { lean_object* x_372; lean_object* x_373; lean_dec(x_369); lean_dec(x_367); -lean_dec(x_366); +lean_dec(x_365); lean_dec(x_2); x_372 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__99; -x_373 = l_Lean_Macro_throwErrorAt___redArg(x_279, x_372, x_368, x_370); +x_373 = l_Lean_Macro_throwErrorAt___redArg(x_279, x_372, x_366, x_370); lean_dec(x_279); return x_373; } @@ -3031,17 +3031,17 @@ if (x_375 == 0) { uint8_t x_376; lean_inc(x_374); -x_376 = l_Lean_Syntax_matchesNull(x_374, x_365); +x_376 = l_Lean_Syntax_matchesNull(x_374, x_364); if (x_376 == 0) { lean_object* x_377; lean_object* x_378; lean_dec(x_374); lean_dec(x_369); lean_dec(x_367); -lean_dec(x_366); +lean_dec(x_365); lean_dec(x_2); x_377 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__99; -x_378 = l_Lean_Macro_throwErrorAt___redArg(x_279, x_377, x_368, x_370); +x_378 = l_Lean_Macro_throwErrorAt___redArg(x_279, x_377, x_366, x_370); lean_dec(x_279); return x_378; } @@ -3052,13 +3052,13 @@ x_379 = l_Lean_Syntax_getArg(x_374, x_64); lean_dec(x_374); x_380 = l_Lean_Syntax_getArg(x_279, x_280); lean_dec(x_279); -x_281 = x_365; -x_282 = x_369; -x_283 = x_380; -x_284 = x_366; -x_285 = x_367; +x_281 = x_364; +x_282 = x_366; +x_283 = x_369; +x_284 = x_365; +x_285 = x_380; x_286 = x_370; -x_287 = x_368; +x_287 = x_367; x_288 = x_379; goto block_363; } @@ -3073,13 +3073,13 @@ x_383 = 0; x_384 = l_Lean_SourceInfo_fromRef(x_279, x_383); lean_dec(x_279); x_385 = l_Lean_Syntax_mkStrLit(x_382, x_384); -x_281 = x_365; -x_282 = x_369; -x_283 = x_381; -x_284 = x_366; -x_285 = x_367; +x_281 = x_364; +x_282 = x_366; +x_283 = x_369; +x_284 = x_365; +x_285 = x_381; x_286 = x_370; -x_287 = x_368; +x_287 = x_367; x_288 = x_385; goto block_363; } @@ -3120,10 +3120,10 @@ goto block_395; } else { -lean_dec_ref(x_400); +lean_dec(x_400); lean_dec(x_399); -lean_dec(x_398); -lean_dec(x_396); +lean_dec_ref(x_398); +lean_dec(x_397); lean_dec(x_279); lean_dec(x_2); return x_401; @@ -3133,41 +3133,41 @@ return x_401; { lean_object* x_412; x_412 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__101; -if (lean_obj_tag(x_405) == 0) -{ -x_364 = x_412; -x_365 = x_406; -x_366 = x_411; -x_367 = x_410; -x_368 = x_409; -x_369 = x_405; -x_370 = x_408; +if (lean_obj_tag(x_408) == 0) +{ +x_364 = x_406; +x_365 = x_407; +x_366 = x_409; +x_367 = x_411; +x_368 = x_412; +x_369 = x_408; +x_370 = x_405; goto block_386; } else { uint8_t x_413; -x_413 = !lean_is_exclusive(x_405); +x_413 = !lean_is_exclusive(x_408); if (x_413 == 0) { lean_object* x_414; lean_object* x_415; uint8_t x_416; -x_414 = lean_ctor_get(x_405, 0); +x_414 = lean_ctor_get(x_408, 0); x_415 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__103; lean_inc(x_414); x_416 = l_Lean_Syntax_isOfKind(x_414, x_415); if (x_416 == 0) { lean_object* x_417; lean_object* x_418; -lean_free_object(x_405); +lean_free_object(x_408); x_417 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_418 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_417, x_409, x_408); +x_418 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_417, x_409, x_405); lean_dec(x_414); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_418; goto block_404; } @@ -3181,7 +3181,7 @@ x_421 = l_Lean_Syntax_isOfKind(x_419, x_420); if (x_421 == 0) { lean_object* x_422; uint8_t x_423; -lean_free_object(x_405); +lean_free_object(x_408); x_422 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__107; lean_inc(x_419); x_423 = l_Lean_Syntax_isOfKind(x_419, x_422); @@ -3191,13 +3191,13 @@ lean_object* x_424; lean_object* x_425; lean_dec(x_419); x_424 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_425 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_424, x_409, x_408); +x_425 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_424, x_409, x_405); lean_dec(x_414); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_425; goto block_404; } @@ -3230,13 +3230,13 @@ x_438 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5_ lean_inc(x_431); x_439 = l_Lean_Syntax_node1(x_431, x_438, x_429); x_440 = l_Lean_Syntax_node2(x_431, x_432, x_437, x_439); -x_387 = x_412; -x_388 = x_406; -x_389 = x_411; -x_390 = x_410; -x_391 = x_409; +x_387 = x_406; +x_388 = x_407; +x_389 = x_409; +x_390 = x_411; +x_391 = x_412; x_392 = x_440; -x_393 = x_408; +x_393 = x_405; goto block_395; } } @@ -3259,16 +3259,16 @@ lean_dec(x_443); lean_dec(x_442); lean_dec(x_441); lean_dec(x_419); -lean_free_object(x_405); +lean_free_object(x_408); x_446 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_447 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_446, x_409, x_408); +x_447 = l_Lean_Macro_throwErrorAt___redArg(x_414, x_446, x_409, x_405); lean_dec(x_414); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_447; goto block_404; } @@ -3278,17 +3278,17 @@ lean_object* x_448; lean_object* x_449; lean_object* x_450; x_448 = l_Lean_Syntax_getArg(x_443, x_280); lean_dec(x_443); x_449 = lean_box(0); -lean_ctor_set(x_405, 0, x_448); +lean_ctor_set(x_408, 0, x_448); lean_inc_ref(x_409); -x_450 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_442, x_441, x_419, x_407, x_406, x_414, x_280, x_449, x_405, x_409, x_408); +x_450 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_442, x_441, x_419, x_410, x_406, x_414, x_280, x_449, x_408, x_409, x_405); lean_dec(x_414); lean_dec(x_419); lean_dec(x_441); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_450; goto block_404; } @@ -3297,19 +3297,19 @@ else { lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_dec(x_443); -lean_free_object(x_405); +lean_free_object(x_408); x_451 = lean_box(0); x_452 = lean_box(0); lean_inc_ref(x_409); -x_453 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_442, x_441, x_419, x_407, x_406, x_414, x_280, x_451, x_452, x_409, x_408); +x_453 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_442, x_441, x_419, x_410, x_406, x_414, x_280, x_451, x_452, x_409, x_405); lean_dec(x_414); lean_dec(x_419); lean_dec(x_441); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_453; goto block_404; } @@ -3319,9 +3319,9 @@ goto block_404; else { lean_object* x_454; lean_object* x_455; uint8_t x_456; -x_454 = lean_ctor_get(x_405, 0); +x_454 = lean_ctor_get(x_408, 0); lean_inc(x_454); -lean_dec(x_405); +lean_dec(x_408); x_455 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__103; lean_inc(x_454); x_456 = l_Lean_Syntax_isOfKind(x_454, x_455); @@ -3330,13 +3330,13 @@ if (x_456 == 0) lean_object* x_457; lean_object* x_458; x_457 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_458 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_457, x_409, x_408); +x_458 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_457, x_409, x_405); lean_dec(x_454); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_458; goto block_404; } @@ -3359,13 +3359,13 @@ lean_object* x_464; lean_object* x_465; lean_dec(x_459); x_464 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_465 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_464, x_409, x_408); +x_465 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_464, x_409, x_405); lean_dec(x_454); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_465; goto block_404; } @@ -3398,13 +3398,13 @@ x_478 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___redArg___lam__5_ lean_inc(x_471); x_479 = l_Lean_Syntax_node1(x_471, x_478, x_469); x_480 = l_Lean_Syntax_node2(x_471, x_472, x_477, x_479); -x_387 = x_412; -x_388 = x_406; -x_389 = x_411; -x_390 = x_410; -x_391 = x_409; +x_387 = x_406; +x_388 = x_407; +x_389 = x_409; +x_390 = x_411; +x_391 = x_412; x_392 = x_480; -x_393 = x_408; +x_393 = x_405; goto block_395; } } @@ -3429,13 +3429,13 @@ lean_dec(x_481); lean_dec(x_459); x_486 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0___closed__5; lean_inc_ref(x_409); -x_487 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_486, x_409, x_408); +x_487 = l_Lean_Macro_throwErrorAt___redArg(x_454, x_486, x_409, x_405); lean_dec(x_454); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_487; goto block_404; } @@ -3448,15 +3448,15 @@ x_489 = lean_box(0); x_490 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_490, 0, x_488); lean_inc_ref(x_409); -x_491 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_482, x_481, x_459, x_407, x_406, x_454, x_280, x_489, x_490, x_409, x_408); +x_491 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_482, x_481, x_459, x_410, x_406, x_454, x_280, x_489, x_490, x_409, x_405); lean_dec(x_454); lean_dec(x_459); lean_dec(x_481); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_491; goto block_404; } @@ -3468,15 +3468,15 @@ lean_dec(x_483); x_492 = lean_box(0); x_493 = lean_box(0); lean_inc_ref(x_409); -x_494 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_482, x_481, x_459, x_407, x_406, x_454, x_280, x_492, x_493, x_409, x_408); +x_494 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___lam__0(x_59, x_482, x_481, x_459, x_410, x_406, x_454, x_280, x_492, x_493, x_409, x_405); lean_dec(x_454); lean_dec(x_459); lean_dec(x_481); -x_396 = x_412; -x_397 = x_406; -x_398 = x_411; -x_399 = x_410; -x_400 = x_409; +x_396 = x_406; +x_397 = x_407; +x_398 = x_409; +x_399 = x_411; +x_400 = x_412; x_401 = x_494; goto block_404; } @@ -3505,7 +3505,7 @@ lean_dec(x_498); lean_dec(x_279); lean_dec(x_2); x_505 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__7; -x_506 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_505, x_499, x_497); +x_506 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_505, x_499, x_496); lean_dec(x_1); return x_506; } @@ -3526,7 +3526,7 @@ lean_dec(x_498); lean_dec(x_279); lean_dec(x_2); x_510 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__7; -x_511 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_510, x_499, x_497); +x_511 = l_Lean_Macro_throwErrorAt___redArg(x_1, x_510, x_499, x_496); lean_dec(x_1); return x_511; } @@ -3538,12 +3538,12 @@ x_512 = l_Lean_Syntax_getArg(x_507, x_280); lean_dec(x_507); x_513 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_513, 0, x_512); -x_405 = x_500; -x_406 = x_496; -x_407 = x_501; -x_408 = x_497; +x_405 = x_496; +x_406 = x_497; +x_407 = x_498; +x_408 = x_500; x_409 = x_499; -x_410 = x_498; +x_410 = x_501; x_411 = x_513; goto block_495; } @@ -3555,12 +3555,12 @@ lean_object* x_514; lean_dec(x_502); lean_dec(x_1); x_514 = lean_box(0); -x_405 = x_500; -x_406 = x_496; -x_407 = x_501; -x_408 = x_497; +x_405 = x_496; +x_406 = x_497; +x_407 = x_498; +x_408 = x_500; x_409 = x_499; -x_410 = x_498; +x_410 = x_501; x_411 = x_514; goto block_495; } @@ -3615,8 +3615,8 @@ x_530 = l_Lean_Syntax_getArg(x_525, x_280); lean_dec(x_525); x_531 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_531, 0, x_530); -x_496 = x_519; -x_497 = x_518; +x_496 = x_518; +x_497 = x_519; x_498 = x_516; x_499 = x_517; x_500 = x_531; @@ -3629,8 +3629,8 @@ else lean_object* x_532; lean_dec(x_520); x_532 = lean_box(0); -x_496 = x_519; -x_497 = x_518; +x_496 = x_518; +x_497 = x_519; x_498 = x_516; x_499 = x_517; x_500 = x_532; @@ -3641,56 +3641,56 @@ goto block_515; block_58: { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_inc(x_19); lean_inc(x_22); -lean_inc(x_24); -x_33 = l_Lean_Syntax_node3(x_24, x_5, x_22, x_19, x_32); -lean_inc_n(x_19, 2); -lean_inc(x_16); -lean_inc(x_24); -x_34 = l_Lean_Syntax_node3(x_24, x_16, x_19, x_19, x_33); -lean_inc(x_24); -x_35 = l_Lean_Syntax_node2(x_24, x_14, x_11, x_34); +lean_inc(x_18); +lean_inc(x_19); +x_33 = l_Lean_Syntax_node3(x_19, x_28, x_18, x_22, x_32); +lean_inc_n(x_22, 2); +lean_inc(x_11); +lean_inc(x_19); +x_34 = l_Lean_Syntax_node3(x_19, x_11, x_22, x_22, x_33); +lean_inc(x_19); +x_35 = l_Lean_Syntax_node2(x_19, x_9, x_20, x_34); x_36 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__0; -x_37 = lean_array_push(x_36, x_15); -lean_inc(x_23); -x_38 = lean_array_push(x_37, x_23); -x_39 = lean_array_push(x_38, x_18); -lean_inc(x_23); -x_40 = lean_array_push(x_39, x_23); -x_41 = lean_array_push(x_40, x_8); -lean_inc(x_23); -x_42 = lean_array_push(x_41, x_23); -x_43 = lean_array_push(x_42, x_25); -lean_inc(x_23); -x_44 = lean_array_push(x_43, x_23); +x_37 = lean_array_push(x_36, x_24); +lean_inc(x_13); +x_38 = lean_array_push(x_37, x_13); +x_39 = lean_array_push(x_38, x_31); +lean_inc(x_13); +x_40 = lean_array_push(x_39, x_13); +x_41 = lean_array_push(x_40, x_25); +lean_inc(x_13); +x_42 = lean_array_push(x_41, x_13); +x_43 = lean_array_push(x_42, x_21); +lean_inc(x_13); +x_44 = lean_array_push(x_43, x_13); x_45 = lean_array_push(x_44, x_35); -x_46 = lean_array_push(x_45, x_23); -lean_inc(x_24); +x_46 = lean_array_push(x_45, x_13); +lean_inc(x_19); x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_24); -lean_ctor_set(x_47, 1, x_16); +lean_ctor_set(x_47, 0, x_19); +lean_ctor_set(x_47, 1, x_11); lean_ctor_set(x_47, 2, x_46); -lean_inc(x_24); -x_48 = l_Lean_Syntax_node1(x_24, x_29, x_47); -lean_inc_n(x_19, 2); -lean_inc(x_24); -x_49 = l_Lean_Syntax_node6(x_24, x_9, x_13, x_19, x_48, x_28, x_19, x_6); +lean_inc(x_19); +x_48 = l_Lean_Syntax_node1(x_19, x_8, x_47); +lean_inc_n(x_22, 2); +lean_inc(x_19); +x_49 = l_Lean_Syntax_node6(x_19, x_6, x_26, x_22, x_48, x_23, x_22, x_15); x_50 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__1; x_51 = l___private_Lake_DSL_Require_0__Lake_DSL_expandDepSpec___closed__2; -x_52 = l_Lean_Name_mkStr4(x_31, x_10, x_50, x_51); -lean_inc_n(x_19, 2); -lean_inc(x_24); -x_53 = l_Lean_Syntax_node2(x_24, x_52, x_19, x_19); +x_52 = l_Lean_Name_mkStr4(x_10, x_7, x_50, x_51); +lean_inc_n(x_22, 2); lean_inc(x_19); -lean_inc(x_24); -x_54 = l_Lean_Syntax_node4(x_24, x_7, x_22, x_49, x_53, x_19); -lean_inc(x_24); -x_55 = l_Lean_Syntax_node5(x_24, x_17, x_27, x_21, x_30, x_54, x_19); -x_56 = l_Lean_Syntax_node2(x_24, x_12, x_26, x_55); +x_53 = l_Lean_Syntax_node2(x_19, x_52, x_22, x_22); +lean_inc(x_22); +lean_inc(x_19); +x_54 = l_Lean_Syntax_node4(x_19, x_27, x_18, x_49, x_53, x_22); +lean_inc(x_19); +x_55 = l_Lean_Syntax_node5(x_19, x_29, x_14, x_17, x_16, x_54, x_22); +x_56 = l_Lean_Syntax_node2(x_19, x_12, x_5, x_55); x_57 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_20); +lean_ctor_set(x_57, 1, x_30); return x_57; } } diff --git a/stage0/stdlib/Lake/DSL/Targets.c b/stage0/stdlib/Lake/DSL/Targets.c index b99bb31a5992..9f7973391b2f 100644 --- a/stage0/stdlib/Lake/DSL/Targets.c +++ b/stage0/stdlib/Lake/DSL/Targets.c @@ -96,7 +96,6 @@ lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__17; static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__13; static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl__1___closed__3; -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_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_object*); LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__24; extern lean_object* l_Lake_Package_keyword; @@ -209,7 +208,6 @@ static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDe lean_object* l_instMonadEST(lean_object*, lean_object*); lean_object* l_Array_empty(lean_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__50; -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_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_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__8; static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabInputDirCommand___closed__2; LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_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*); @@ -300,7 +298,6 @@ uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_elabCommand___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand__1___closed__0; static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___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_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0_spec__0_spec__0___redArg(lean_object*, lean_object*); static lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__8; static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; @@ -520,7 +517,6 @@ static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_ static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0_spec__0_spec__0___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_expandPackageFacetDecl___closed__8; -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___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_Lake_DSL_mkModuleFacetDecl___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) { _start: { @@ -5004,7 +5000,7 @@ return x_31; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_163; uint8_t x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_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_371; lean_object* x_372; lean_object* x_373; lean_object* x_387; lean_object* x_394; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_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_371; lean_object* x_372; lean_object* x_373; lean_object* x_387; lean_object* x_394; x_32 = lean_unsigned_to_nat(0u); x_33 = l_Lean_Syntax_getArg(x_1, x_32); x_34 = lean_unsigned_to_nat(1u); @@ -5048,180 +5044,180 @@ goto block_393; block_162: { lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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_inc_ref(x_38); -x_64 = l_Array_append___redArg(x_38, x_63); +lean_inc_ref(x_43); +x_64 = l_Array_append___redArg(x_43, x_63); lean_dec_ref(x_63); -lean_inc(x_55); -lean_inc(x_39); +lean_inc(x_56); +lean_inc(x_42); x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_39); -lean_ctor_set(x_65, 1, x_55); +lean_ctor_set(x_65, 0, x_42); +lean_ctor_set(x_65, 1, x_56); lean_ctor_set(x_65, 2, x_64); x_66 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_61); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_67 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_66); +lean_inc_ref(x_51); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_67 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_66); x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_39); +lean_inc(x_42); x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_39); +lean_ctor_set(x_69, 0, x_42); lean_ctor_set(x_69, 1, x_68); x_70 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -x_71 = l_Lean_Syntax_SepArray_ofElems(x_70, x_43); -lean_dec_ref(x_43); -lean_inc_ref(x_38); -x_72 = l_Array_append___redArg(x_38, x_71); +x_71 = l_Lean_Syntax_SepArray_ofElems(x_70, x_60); +lean_dec_ref(x_60); +lean_inc_ref(x_43); +x_72 = l_Array_append___redArg(x_43, x_71); lean_dec_ref(x_71); -lean_inc(x_55); -lean_inc(x_39); +lean_inc(x_56); +lean_inc(x_42); x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_55); +lean_ctor_set(x_73, 0, x_42); +lean_ctor_set(x_73, 1, x_56); lean_ctor_set(x_73, 2, x_72); x_74 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_39); +lean_inc(x_42); x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_39); +lean_ctor_set(x_75, 0, x_42); lean_ctor_set(x_75, 1, x_74); -lean_inc(x_39); -x_76 = l_Lean_Syntax_node3(x_39, x_67, x_69, x_73, x_75); -lean_inc(x_55); -lean_inc(x_39); -x_77 = l_Lean_Syntax_node1(x_39, x_55, x_76); -lean_inc_n(x_51, 5); -lean_inc(x_39); -x_78 = l_Lean_Syntax_node7(x_39, x_47, x_65, x_77, x_51, x_51, x_51, x_51, x_51); +lean_inc(x_42); +x_76 = l_Lean_Syntax_node3(x_42, x_67, x_69, x_73, x_75); +lean_inc(x_56); +lean_inc(x_42); +x_77 = l_Lean_Syntax_node1(x_42, x_56, x_76); +lean_inc_n(x_44, 5); +lean_inc(x_42); +x_78 = l_Lean_Syntax_node7(x_42, x_39, x_65, x_77, x_44, x_44, x_44, x_44, x_44); x_79 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; -lean_inc_ref(x_42); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_80 = l_Lean_Name_mkStr4(x_58, x_57, x_42, x_79); -lean_inc(x_39); +lean_inc_ref(x_48); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_80 = l_Lean_Name_mkStr4(x_59, x_53, x_48, x_79); +lean_inc(x_42); x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_39); +lean_ctor_set(x_81, 0, x_42); lean_ctor_set(x_81, 1, x_79); x_82 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; -lean_inc_ref(x_42); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_83 = l_Lean_Name_mkStr4(x_58, x_57, x_42, x_82); +lean_inc_ref(x_48); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_83 = l_Lean_Name_mkStr4(x_59, x_53, x_48, x_82); x_84 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_85 = lean_box(2); -lean_inc(x_55); +lean_inc(x_56); x_86 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_55); +lean_ctor_set(x_86, 1, x_56); lean_ctor_set(x_86, 2, x_84); x_87 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_88 = lean_array_push(x_87, x_50); +x_88 = lean_array_push(x_87, x_37); x_89 = lean_array_push(x_88, x_86); x_90 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_90, 0, x_85); lean_ctor_set(x_90, 1, x_83); lean_ctor_set(x_90, 2, x_89); x_91 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_42); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_92 = l_Lean_Name_mkStr4(x_58, x_57, x_42, x_91); +lean_inc_ref(x_48); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_92 = l_Lean_Name_mkStr4(x_59, x_53, x_48, x_91); x_93 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__4; x_94 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__5; -lean_inc(x_54); -lean_inc(x_53); -x_95 = l_Lean_addMacroScope(x_53, x_94, x_54); +lean_inc(x_58); +lean_inc(x_55); +x_95 = l_Lean_addMacroScope(x_55, x_94, x_58); x_96 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__7; -lean_inc(x_62); +lean_inc(x_49); x_97 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_62); -lean_inc(x_39); +lean_ctor_set(x_97, 1, x_49); +lean_inc(x_42); x_98 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_98, 0, x_39); +lean_ctor_set(x_98, 0, x_42); lean_ctor_set(x_98, 1, x_93); lean_ctor_set(x_98, 2, x_95); lean_ctor_set(x_98, 3, x_97); -lean_inc(x_39); -x_99 = l_Lean_Syntax_node2(x_39, x_37, x_46, x_98); -lean_inc(x_55); -lean_inc(x_39); -x_100 = l_Lean_Syntax_node1(x_39, x_55, x_99); -lean_inc(x_51); -lean_inc(x_39); -x_101 = l_Lean_Syntax_node2(x_39, x_92, x_51, x_100); +lean_inc(x_42); +x_99 = l_Lean_Syntax_node2(x_42, x_61, x_50, x_98); +lean_inc(x_56); +lean_inc(x_42); +x_100 = l_Lean_Syntax_node1(x_42, x_56, x_99); +lean_inc(x_44); +lean_inc(x_42); +x_101 = l_Lean_Syntax_node2(x_42, x_92, x_44, x_100); x_102 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_39); +lean_inc(x_42); x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_39); +lean_ctor_set(x_103, 0, x_42); lean_ctor_set(x_103, 1, x_102); x_104 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_61); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_105 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_104); +lean_inc_ref(x_51); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_105 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_104); x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__9; x_107 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__11; -lean_inc(x_54); -lean_inc(x_53); -x_108 = l_Lean_addMacroScope(x_53, x_107, x_54); +lean_inc(x_58); +lean_inc(x_55); +x_108 = l_Lean_addMacroScope(x_55, x_107, x_58); x_109 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__12; -lean_inc(x_62); +lean_inc(x_49); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_62); -lean_inc(x_39); +lean_ctor_set(x_110, 1, x_49); +lean_inc(x_42); x_111 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_111, 0, x_39); +lean_ctor_set(x_111, 0, x_42); lean_ctor_set(x_111, 1, x_106); lean_ctor_set(x_111, 2, x_108); lean_ctor_set(x_111, 3, x_110); x_112 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__21; -lean_inc_ref(x_61); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_113 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_112); +lean_inc_ref(x_51); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_113 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_112); x_114 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_61); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_115 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_114); +lean_inc_ref(x_51); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_115 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_114); x_116 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_39); +lean_inc(x_42); x_117 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_117, 0, x_39); +lean_ctor_set(x_117, 0, x_42); lean_ctor_set(x_117, 1, x_116); x_118 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_119 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_120 = lean_box(0); -x_121 = l_Lean_addMacroScope(x_53, x_120, x_54); +x_121 = l_Lean_addMacroScope(x_55, x_120, x_58); x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; -lean_inc_ref(x_42); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_123 = l_Lean_Name_mkStr3(x_58, x_57, x_42); +lean_inc_ref(x_48); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_123 = l_Lean_Name_mkStr3(x_59, x_53, x_48); x_124 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_124, 0, x_123); x_125 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_58); -x_126 = l_Lean_Name_mkStr3(x_58, x_125, x_42); +lean_inc_ref(x_59); +x_126 = l_Lean_Name_mkStr3(x_59, x_125, x_48); x_127 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_127, 0, x_126); -lean_inc_ref(x_58); -x_128 = l_Lean_Name_mkStr2(x_58, x_125); +lean_inc_ref(x_59); +x_128 = l_Lean_Name_mkStr2(x_59, x_125); x_129 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_129, 0, x_128); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_130 = l_Lean_Name_mkStr2(x_58, x_57); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_130 = l_Lean_Name_mkStr2(x_59, x_53); x_131 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_131, 0, x_130); -lean_inc_ref(x_58); -x_132 = l_Lean_Name_mkStr1(x_58); +lean_inc_ref(x_59); +x_132 = l_Lean_Name_mkStr1(x_59); x_133 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_133, 0, x_132); x_134 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_62); +lean_ctor_set(x_134, 1, x_49); x_135 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_135, 0, x_131); lean_ctor_set(x_135, 1, x_134); @@ -5237,99 +5233,99 @@ lean_ctor_set(x_138, 1, x_137); x_139 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_139, 0, x_122); lean_ctor_set(x_139, 1, x_138); -lean_inc(x_39); +lean_inc(x_42); x_140 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_140, 0, x_39); +lean_ctor_set(x_140, 0, x_42); lean_ctor_set(x_140, 1, x_119); lean_ctor_set(x_140, 2, x_121); lean_ctor_set(x_140, 3, x_139); -lean_inc(x_39); -x_141 = l_Lean_Syntax_node1(x_39, x_118, x_140); -lean_inc(x_39); -x_142 = l_Lean_Syntax_node2(x_39, x_115, x_117, x_141); +lean_inc(x_42); +x_141 = l_Lean_Syntax_node1(x_42, x_118, x_140); +lean_inc(x_42); +x_142 = l_Lean_Syntax_node2(x_42, x_115, x_117, x_141); x_143 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__31; -lean_inc_ref(x_61); -lean_inc_ref(x_57); -lean_inc_ref(x_58); -x_144 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_143); -lean_inc(x_39); +lean_inc_ref(x_51); +lean_inc_ref(x_53); +lean_inc_ref(x_59); +x_144 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_143); +lean_inc(x_42); x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_39); +lean_ctor_set(x_145, 0, x_42); lean_ctor_set(x_145, 1, x_143); x_146 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__32; -x_147 = l_Lean_Name_mkStr4(x_58, x_57, x_61, x_146); -lean_inc(x_55); -lean_inc(x_39); -x_148 = l_Lean_Syntax_node1(x_39, x_55, x_48); +x_147 = l_Lean_Name_mkStr4(x_59, x_53, x_51, x_146); +lean_inc(x_56); +lean_inc(x_42); +x_148 = l_Lean_Syntax_node1(x_42, x_56, x_57); x_149 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__33; -lean_inc(x_39); +lean_inc(x_42); x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_39); +lean_ctor_set(x_150, 0, x_42); lean_ctor_set(x_150, 1, x_149); -lean_inc(x_51); -lean_inc(x_39); -x_151 = l_Lean_Syntax_node4(x_39, x_147, x_148, x_51, x_150, x_40); -lean_inc(x_39); -x_152 = l_Lean_Syntax_node2(x_39, x_144, x_145, x_151); +lean_inc(x_44); +lean_inc(x_42); +x_151 = l_Lean_Syntax_node4(x_42, x_147, x_148, x_44, x_150, x_38); +lean_inc(x_42); +x_152 = l_Lean_Syntax_node2(x_42, x_144, x_145, x_151); x_153 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_39); +lean_inc(x_42); x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_39); +lean_ctor_set(x_154, 0, x_42); lean_ctor_set(x_154, 1, x_153); -lean_inc(x_39); -x_155 = l_Lean_Syntax_node3(x_39, x_113, x_142, x_152, x_154); -lean_inc(x_55); -lean_inc(x_39); -x_156 = l_Lean_Syntax_node3(x_39, x_55, x_56, x_44, x_155); -lean_inc(x_39); -x_157 = l_Lean_Syntax_node2(x_39, x_105, x_111, x_156); -lean_inc(x_51); -lean_inc(x_39); -x_158 = l_Lean_Syntax_node2(x_39, x_52, x_51, x_51); -if (lean_obj_tag(x_41) == 0) +lean_inc(x_42); +x_155 = l_Lean_Syntax_node3(x_42, x_113, x_142, x_152, x_154); +lean_inc(x_56); +lean_inc(x_42); +x_156 = l_Lean_Syntax_node3(x_42, x_56, x_40, x_45, x_155); +lean_inc(x_42); +x_157 = l_Lean_Syntax_node2(x_42, x_105, x_111, x_156); +lean_inc(x_44); +lean_inc(x_42); +x_158 = l_Lean_Syntax_node2(x_42, x_54, x_44, x_44); +if (lean_obj_tag(x_46) == 0) { lean_object* x_159; x_159 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_4 = x_38; -x_5 = x_39; -x_6 = x_55; -x_7 = x_157; -x_8 = x_81; -x_9 = x_80; -x_10 = x_59; -x_11 = x_60; -x_12 = x_45; -x_13 = x_103; -x_14 = x_90; -x_15 = x_78; -x_16 = x_158; -x_17 = x_49; -x_18 = x_101; +x_4 = x_101; +x_5 = x_41; +x_6 = x_80; +x_7 = x_42; +x_8 = x_78; +x_9 = x_43; +x_10 = x_90; +x_11 = x_103; +x_12 = x_158; +x_13 = x_52; +x_14 = x_157; +x_15 = x_81; +x_16 = x_56; +x_17 = x_62; +x_18 = x_47; x_19 = x_159; goto block_27; } else { lean_object* x_160; lean_object* x_161; -x_160 = lean_ctor_get(x_41, 0); +x_160 = lean_ctor_get(x_46, 0); lean_inc(x_160); -lean_dec_ref(x_41); +lean_dec_ref(x_46); x_161 = l_Array_mkArray1___redArg(x_160); -x_4 = x_38; -x_5 = x_39; -x_6 = x_55; -x_7 = x_157; -x_8 = x_81; -x_9 = x_80; -x_10 = x_59; -x_11 = x_60; -x_12 = x_45; -x_13 = x_103; -x_14 = x_90; -x_15 = x_78; -x_16 = x_158; -x_17 = x_49; -x_18 = x_101; +x_4 = x_101; +x_5 = x_41; +x_6 = x_80; +x_7 = x_42; +x_8 = x_78; +x_9 = x_43; +x_10 = x_90; +x_11 = x_103; +x_12 = x_158; +x_13 = x_52; +x_14 = x_157; +x_15 = x_81; +x_16 = x_56; +x_17 = x_62; +x_18 = x_47; x_19 = x_161; goto block_27; } @@ -5337,7 +5333,7 @@ goto block_27; block_210: { lean_object* x_189; -x_189 = l_Lake_DSL_expandOptSimpleBinder(x_187, x_181, x_182); +x_189 = l_Lake_DSL_expandOptSimpleBinder(x_168, x_172, x_165); if (lean_obj_tag(x_189) == 0) { lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; @@ -5346,7 +5342,7 @@ lean_inc(x_190); x_191 = lean_ctor_get(x_189, 1); lean_inc(x_191); lean_dec_ref(x_189); -x_192 = l_Lean_mkIdentFrom(x_172, x_188, x_164); +x_192 = l_Lean_mkIdentFrom(x_176, x_188, x_184); x_193 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__14; x_194 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__15; lean_inc(x_166); @@ -5358,87 +5354,87 @@ lean_inc(x_166); x_197 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_197, 0, x_166); lean_ctor_set(x_197, 1, x_196); -lean_inc(x_176); +lean_inc(x_164); lean_inc_ref(x_197); -lean_inc(x_173); +lean_inc(x_169); lean_inc(x_166); -x_198 = l_Lean_Syntax_node5(x_166, x_193, x_173, x_195, x_172, x_197, x_176); +x_198 = l_Lean_Syntax_node5(x_166, x_193, x_169, x_195, x_176, x_197, x_164); x_199 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_169); -lean_inc_ref(x_179); -lean_inc_ref(x_180); -x_200 = l_Lean_Name_mkStr4(x_180, x_179, x_169, x_199); +lean_inc_ref(x_174); +lean_inc_ref(x_178); +lean_inc_ref(x_183); +x_200 = l_Lean_Name_mkStr4(x_183, x_178, x_174, x_199); x_201 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_169); -lean_inc_ref(x_179); -lean_inc_ref(x_180); -x_202 = l_Lean_Name_mkStr4(x_180, x_179, x_169, x_201); -if (lean_obj_tag(x_186) == 0) +lean_inc_ref(x_174); +lean_inc_ref(x_178); +lean_inc_ref(x_183); +x_202 = l_Lean_Name_mkStr4(x_183, x_178, x_174, x_201); +if (lean_obj_tag(x_187) == 0) { lean_object* x_203; x_203 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_37 = x_163; -x_38 = x_165; -x_39 = x_166; -x_40 = x_167; -x_41 = x_168; -x_42 = x_169; -x_43 = x_170; -x_44 = x_171; -x_45 = x_198; -x_46 = x_197; -x_47 = x_202; -x_48 = x_190; -x_49 = x_191; -x_50 = x_192; -x_51 = x_173; -x_52 = x_174; -x_53 = x_175; -x_54 = x_177; -x_55 = x_178; -x_56 = x_176; -x_57 = x_179; -x_58 = x_180; +x_37 = x_192; +x_38 = x_163; +x_39 = x_202; +x_40 = x_164; +x_41 = x_198; +x_42 = x_166; +x_43 = x_167; +x_44 = x_169; +x_45 = x_170; +x_46 = x_171; +x_47 = x_173; +x_48 = x_174; +x_49 = x_175; +x_50 = x_197; +x_51 = x_177; +x_52 = x_191; +x_53 = x_178; +x_54 = x_179; +x_55 = x_180; +x_56 = x_181; +x_57 = x_190; +x_58 = x_182; x_59 = x_183; -x_60 = x_200; -x_61 = x_184; -x_62 = x_185; +x_60 = x_185; +x_61 = x_186; +x_62 = x_200; x_63 = x_203; goto block_162; } else { lean_object* x_204; lean_object* x_205; -x_204 = lean_ctor_get(x_186, 0); +x_204 = lean_ctor_get(x_187, 0); lean_inc(x_204); -lean_dec_ref(x_186); +lean_dec_ref(x_187); x_205 = l_Array_mkArray1___redArg(x_204); -x_37 = x_163; -x_38 = x_165; -x_39 = x_166; -x_40 = x_167; -x_41 = x_168; -x_42 = x_169; -x_43 = x_170; -x_44 = x_171; -x_45 = x_198; -x_46 = x_197; -x_47 = x_202; -x_48 = x_190; -x_49 = x_191; -x_50 = x_192; -x_51 = x_173; -x_52 = x_174; -x_53 = x_175; -x_54 = x_177; -x_55 = x_178; -x_56 = x_176; -x_57 = x_179; -x_58 = x_180; +x_37 = x_192; +x_38 = x_163; +x_39 = x_202; +x_40 = x_164; +x_41 = x_198; +x_42 = x_166; +x_43 = x_167; +x_44 = x_169; +x_45 = x_170; +x_46 = x_171; +x_47 = x_173; +x_48 = x_174; +x_49 = x_175; +x_50 = x_197; +x_51 = x_177; +x_52 = x_191; +x_53 = x_178; +x_54 = x_179; +x_55 = x_180; +x_56 = x_181; +x_57 = x_190; +x_58 = x_182; x_59 = x_183; -x_60 = x_200; -x_61 = x_184; -x_62 = x_185; +x_60 = x_185; +x_61 = x_186; +x_62 = x_200; x_63 = x_205; goto block_162; } @@ -5447,26 +5443,26 @@ else { uint8_t x_206; lean_dec(x_188); +lean_dec(x_187); lean_dec(x_186); -lean_dec(x_185); -lean_dec_ref(x_184); -lean_dec(x_183); -lean_dec_ref(x_180); -lean_dec_ref(x_179); -lean_dec(x_178); -lean_dec(x_177); +lean_dec_ref(x_185); +lean_dec_ref(x_183); +lean_dec(x_182); +lean_dec(x_181); +lean_dec(x_180); +lean_dec(x_179); +lean_dec_ref(x_178); +lean_dec_ref(x_177); lean_dec(x_176); lean_dec(x_175); -lean_dec(x_174); +lean_dec_ref(x_174); lean_dec(x_173); -lean_dec(x_172); lean_dec(x_171); -lean_dec_ref(x_170); -lean_dec_ref(x_169); -lean_dec(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec(x_169); +lean_dec_ref(x_167); lean_dec(x_166); -lean_dec_ref(x_165); +lean_dec(x_164); lean_dec(x_163); x_206 = !lean_is_exclusive(x_189); if (x_206 == 0) @@ -5509,15 +5505,15 @@ x_233 = 0; x_234 = l_Lean_SourceInfo_fromRef(x_232, x_233); lean_dec(x_232); x_235 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_220); -lean_inc_ref(x_216); lean_inc_ref(x_217); -x_236 = l_Lean_Name_mkStr4(x_217, x_216, x_220, x_235); -x_237 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; +lean_inc_ref(x_218); lean_inc_ref(x_220); -lean_inc_ref(x_216); +x_236 = l_Lean_Name_mkStr4(x_220, x_218, x_217, x_235); +x_237 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; lean_inc_ref(x_217); -x_238 = l_Lean_Name_mkStr4(x_217, x_216, x_220, x_237); +lean_inc_ref(x_218); +lean_inc_ref(x_220); +x_238 = l_Lean_Name_mkStr4(x_220, x_218, x_217, x_237); x_239 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_240 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_234); @@ -5530,9 +5526,9 @@ lean_inc(x_234); x_242 = l_Lean_Syntax_node1(x_234, x_238, x_241); x_243 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_244 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_216); -lean_inc_ref(x_217); -x_245 = l_Lean_Name_mkStr4(x_217, x_216, x_243, x_244); +lean_inc_ref(x_218); +lean_inc_ref(x_220); +x_245 = l_Lean_Name_mkStr4(x_220, x_218, x_243, x_244); x_246 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__17; x_247 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__19; lean_inc(x_230); @@ -5555,7 +5551,7 @@ x_254 = lean_array_push(x_253, x_252); x_255 = l_Lake_DSL_expandAttrs(x_214); x_256 = l_Array_append___redArg(x_254, x_255); lean_dec_ref(x_255); -x_257 = l_Lake_DSL_expandIdentOrStrAsIdent(x_223); +x_257 = l_Lake_DSL_expandIdentOrStrAsIdent(x_215); x_258 = l_Lean_TSyntax_getId(x_257); lean_inc(x_258); lean_inc(x_257); @@ -5566,30 +5562,30 @@ if (x_260 == 0) lean_object* x_261; x_261 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___lam__0(x_258); x_163 = x_212; -x_164 = x_233; -x_165 = x_240; +x_164 = x_213; +x_165 = x_227; x_166 = x_234; -x_167 = x_215; -x_168 = x_225; -x_169 = x_219; -x_170 = x_256; -x_171 = x_259; -x_172 = x_257; -x_173 = x_241; -x_174 = x_224; -x_175 = x_229; -x_176 = x_213; -x_177 = x_230; -x_178 = x_239; -x_179 = x_216; -x_180 = x_217; -x_181 = x_226; -x_182 = x_227; -x_183 = x_218; -x_184 = x_220; -x_185 = x_249; +x_167 = x_240; +x_168 = x_216; +x_169 = x_241; +x_170 = x_259; +x_171 = x_225; +x_172 = x_226; +x_173 = x_222; +x_174 = x_223; +x_175 = x_249; +x_176 = x_257; +x_177 = x_217; +x_178 = x_218; +x_179 = x_219; +x_180 = x_229; +x_181 = x_239; +x_182 = x_230; +x_183 = x_220; +x_184 = x_233; +x_185 = x_256; x_186 = x_221; -x_187 = x_222; +x_187 = x_224; x_188 = x_261; goto block_210; } @@ -5606,30 +5602,30 @@ x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___lam__0 lean_ctor_set(x_262, 0, x_265); x_266 = l_Lean_MacroScopesView_review(x_262); x_163 = x_212; -x_164 = x_233; -x_165 = x_240; +x_164 = x_213; +x_165 = x_227; x_166 = x_234; -x_167 = x_215; -x_168 = x_225; -x_169 = x_219; -x_170 = x_256; -x_171 = x_259; -x_172 = x_257; -x_173 = x_241; -x_174 = x_224; -x_175 = x_229; -x_176 = x_213; -x_177 = x_230; -x_178 = x_239; -x_179 = x_216; -x_180 = x_217; -x_181 = x_226; -x_182 = x_227; -x_183 = x_218; -x_184 = x_220; -x_185 = x_249; +x_167 = x_240; +x_168 = x_216; +x_169 = x_241; +x_170 = x_259; +x_171 = x_225; +x_172 = x_226; +x_173 = x_222; +x_174 = x_223; +x_175 = x_249; +x_176 = x_257; +x_177 = x_217; +x_178 = x_218; +x_179 = x_219; +x_180 = x_229; +x_181 = x_239; +x_182 = x_230; +x_183 = x_220; +x_184 = x_233; +x_185 = x_256; x_186 = x_221; -x_187 = x_222; +x_187 = x_224; x_188 = x_266; goto block_210; } @@ -5653,30 +5649,30 @@ lean_ctor_set(x_272, 2, x_269); lean_ctor_set(x_272, 3, x_270); x_273 = l_Lean_MacroScopesView_review(x_272); x_163 = x_212; -x_164 = x_233; -x_165 = x_240; +x_164 = x_213; +x_165 = x_227; x_166 = x_234; -x_167 = x_215; -x_168 = x_225; -x_169 = x_219; -x_170 = x_256; -x_171 = x_259; -x_172 = x_257; -x_173 = x_241; -x_174 = x_224; -x_175 = x_229; -x_176 = x_213; -x_177 = x_230; -x_178 = x_239; -x_179 = x_216; -x_180 = x_217; -x_181 = x_226; -x_182 = x_227; -x_183 = x_218; -x_184 = x_220; -x_185 = x_249; +x_167 = x_240; +x_168 = x_216; +x_169 = x_241; +x_170 = x_259; +x_171 = x_225; +x_172 = x_226; +x_173 = x_222; +x_174 = x_223; +x_175 = x_249; +x_176 = x_257; +x_177 = x_217; +x_178 = x_218; +x_179 = x_219; +x_180 = x_229; +x_181 = x_239; +x_182 = x_230; +x_183 = x_220; +x_184 = x_233; +x_185 = x_256; x_186 = x_221; -x_187 = x_222; +x_187 = x_224; x_188 = x_273; goto block_210; } @@ -5715,15 +5711,15 @@ x_282 = 0; x_283 = l_Lean_SourceInfo_fromRef(x_280, x_282); lean_dec(x_280); x_284 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_220); -lean_inc_ref(x_216); lean_inc_ref(x_217); -x_285 = l_Lean_Name_mkStr4(x_217, x_216, x_220, x_284); -x_286 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; +lean_inc_ref(x_218); lean_inc_ref(x_220); -lean_inc_ref(x_216); +x_285 = l_Lean_Name_mkStr4(x_220, x_218, x_217, x_284); +x_286 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; lean_inc_ref(x_217); -x_287 = l_Lean_Name_mkStr4(x_217, x_216, x_220, x_286); +lean_inc_ref(x_218); +lean_inc_ref(x_220); +x_287 = l_Lean_Name_mkStr4(x_220, x_218, x_217, x_286); x_288 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_289 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_283); @@ -5736,9 +5732,9 @@ lean_inc(x_283); x_291 = l_Lean_Syntax_node1(x_283, x_287, x_290); x_292 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_293 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_216); -lean_inc_ref(x_217); -x_294 = l_Lean_Name_mkStr4(x_217, x_216, x_292, x_293); +lean_inc_ref(x_218); +lean_inc_ref(x_220); +x_294 = l_Lean_Name_mkStr4(x_220, x_218, x_292, x_293); x_295 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__17; x_296 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___closed__19; lean_inc(x_276); @@ -5761,7 +5757,7 @@ x_303 = lean_array_push(x_302, x_301); x_304 = l_Lake_DSL_expandAttrs(x_214); x_305 = l_Array_append___redArg(x_303, x_304); lean_dec_ref(x_304); -x_306 = l_Lake_DSL_expandIdentOrStrAsIdent(x_223); +x_306 = l_Lake_DSL_expandIdentOrStrAsIdent(x_215); x_307 = l_Lean_TSyntax_getId(x_306); lean_inc(x_307); lean_inc(x_306); @@ -5772,30 +5768,30 @@ if (x_309 == 0) lean_object* x_310; x_310 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandLibraryFacetDecl___lam__0(x_307); x_163 = x_212; -x_164 = x_282; -x_165 = x_289; +x_164 = x_213; +x_165 = x_227; x_166 = x_283; -x_167 = x_215; -x_168 = x_225; -x_169 = x_219; -x_170 = x_305; -x_171 = x_308; -x_172 = x_306; -x_173 = x_290; -x_174 = x_224; -x_175 = x_275; -x_176 = x_213; -x_177 = x_276; -x_178 = x_288; -x_179 = x_216; -x_180 = x_217; -x_181 = x_281; -x_182 = x_227; -x_183 = x_218; -x_184 = x_220; -x_185 = x_298; +x_167 = x_289; +x_168 = x_216; +x_169 = x_290; +x_170 = x_308; +x_171 = x_225; +x_172 = x_281; +x_173 = x_222; +x_174 = x_223; +x_175 = x_298; +x_176 = x_306; +x_177 = x_217; +x_178 = x_218; +x_179 = x_219; +x_180 = x_275; +x_181 = x_288; +x_182 = x_276; +x_183 = x_220; +x_184 = x_282; +x_185 = x_305; x_186 = x_221; -x_187 = x_222; +x_187 = x_224; x_188 = x_310; goto block_210; } @@ -5833,30 +5829,30 @@ lean_ctor_set(x_318, 2, x_314); lean_ctor_set(x_318, 3, x_315); x_319 = l_Lean_MacroScopesView_review(x_318); x_163 = x_212; -x_164 = x_282; -x_165 = x_289; +x_164 = x_213; +x_165 = x_227; x_166 = x_283; -x_167 = x_215; -x_168 = x_225; -x_169 = x_219; -x_170 = x_305; -x_171 = x_308; -x_172 = x_306; -x_173 = x_290; -x_174 = x_224; -x_175 = x_275; -x_176 = x_213; -x_177 = x_276; -x_178 = x_288; -x_179 = x_216; -x_180 = x_217; -x_181 = x_281; -x_182 = x_227; -x_183 = x_218; -x_184 = x_220; -x_185 = x_298; +x_167 = x_289; +x_168 = x_216; +x_169 = x_290; +x_170 = x_308; +x_171 = x_225; +x_172 = x_281; +x_173 = x_222; +x_174 = x_223; +x_175 = x_298; +x_176 = x_306; +x_177 = x_217; +x_178 = x_218; +x_179 = x_219; +x_180 = x_275; +x_181 = x_288; +x_182 = x_276; +x_183 = x_220; +x_184 = x_282; +x_185 = x_305; x_186 = x_221; -x_187 = x_222; +x_187 = x_224; x_188 = x_319; goto block_210; } @@ -6034,19 +6030,19 @@ lean_object* x_368; lean_dec(x_322); x_368 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_368, 0, x_363); -x_212 = x_333; +x_212 = x_357; x_213 = x_356; -x_214 = x_324; -x_215 = x_357; -x_216 = x_331; -x_217 = x_330; -x_218 = x_339; -x_219 = x_338; -x_220 = x_332; -x_221 = x_323; -x_222 = x_326; -x_223 = x_325; -x_224 = x_344; +x_214 = x_325; +x_215 = x_323; +x_216 = x_326; +x_217 = x_332; +x_218 = x_331; +x_219 = x_344; +x_220 = x_330; +x_221 = x_333; +x_222 = x_339; +x_223 = x_338; +x_224 = x_324; x_225 = x_368; x_226 = x_327; x_227 = x_328; @@ -6060,19 +6056,19 @@ lean_object* x_369; lean_dec(x_358); lean_dec(x_322); x_369 = lean_box(0); -x_212 = x_333; +x_212 = x_357; x_213 = x_356; -x_214 = x_324; -x_215 = x_357; -x_216 = x_331; -x_217 = x_330; -x_218 = x_339; -x_219 = x_338; -x_220 = x_332; -x_221 = x_323; -x_222 = x_326; -x_223 = x_325; -x_224 = x_344; +x_214 = x_325; +x_215 = x_323; +x_216 = x_326; +x_217 = x_332; +x_218 = x_331; +x_219 = x_344; +x_220 = x_330; +x_221 = x_333; +x_222 = x_339; +x_223 = x_338; +x_224 = x_324; x_225 = x_369; x_226 = x_327; x_227 = x_328; @@ -6131,9 +6127,9 @@ x_383 = l_Lean_Syntax_getArg(x_378, x_32); lean_dec(x_378); x_384 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_384, 0, x_383); -x_323 = x_373; -x_324 = x_372; -x_325 = x_377; +x_323 = x_377; +x_324 = x_373; +x_325 = x_372; x_326 = x_384; x_327 = x_2; x_328 = x_3; @@ -6145,9 +6141,9 @@ else lean_object* x_385; lean_dec(x_378); x_385 = lean_box(0); -x_323 = x_373; -x_324 = x_372; -x_325 = x_377; +x_323 = x_377; +x_324 = x_373; +x_325 = x_372; x_326 = x_385; x_327 = x_2; x_328 = x_3; @@ -6196,24 +6192,24 @@ goto block_386; block_27: { lean_object* x_20; 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 = l_Array_append___redArg(x_4, x_19); +x_20 = l_Array_append___redArg(x_9, x_19); lean_dec_ref(x_19); -lean_inc(x_6); -lean_inc(x_5); +lean_inc(x_16); +lean_inc(x_7); x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_5); -lean_ctor_set(x_21, 1, x_6); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_16); lean_ctor_set(x_21, 2, x_20); -lean_inc(x_5); -x_22 = l_Lean_Syntax_node4(x_5, x_10, x_13, x_7, x_16, x_21); -lean_inc(x_5); -x_23 = l_Lean_Syntax_node4(x_5, x_9, x_8, x_14, x_18, x_22); -lean_inc(x_5); -x_24 = l_Lean_Syntax_node2(x_5, x_11, x_15, x_23); -x_25 = l_Lean_Syntax_node2(x_5, x_6, x_12, x_24); +lean_inc(x_7); +x_22 = l_Lean_Syntax_node4(x_7, x_18, x_11, x_14, x_12, x_21); +lean_inc(x_7); +x_23 = l_Lean_Syntax_node4(x_7, x_6, x_15, x_10, x_4, x_22); +lean_inc(x_7); +x_24 = l_Lean_Syntax_node2(x_7, x_17, x_8, x_23); +x_25 = l_Lean_Syntax_node2(x_7, x_16, x_5, 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_17); +lean_ctor_set(x_26, 1, x_13); return x_26; } } @@ -6905,40 +6901,40 @@ lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_inc_ref(x_118); x_132 = l_Array_append___redArg(x_118, x_131); lean_dec_ref(x_131); -lean_inc(x_117); -lean_inc(x_120); +lean_inc(x_109); +lean_inc(x_106); x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_120); -lean_ctor_set(x_133, 1, x_117); +lean_ctor_set(x_133, 0, x_106); +lean_ctor_set(x_133, 1, x_109); lean_ctor_set(x_133, 2, x_132); -lean_inc_n(x_113, 6); -lean_inc(x_111); -lean_inc(x_120); -x_134 = l_Lean_Syntax_node7(x_120, x_111, x_133, x_113, x_113, x_113, x_113, x_113, x_113); +lean_inc_n(x_119, 6); +lean_inc(x_127); +lean_inc(x_106); +x_134 = l_Lean_Syntax_node7(x_106, x_127, x_133, x_119, x_119, x_119, x_119, x_119, x_119); x_135 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; -lean_inc_ref(x_123); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_136 = l_Lean_Name_mkStr4(x_99, x_103, x_123, x_135); -lean_inc(x_120); +lean_inc_ref(x_101); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_136 = l_Lean_Name_mkStr4(x_108, x_105, x_101, x_135); +lean_inc(x_106); x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_120); +lean_ctor_set(x_137, 0, x_106); lean_ctor_set(x_137, 1, x_135); x_138 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; -lean_inc_ref(x_123); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_139 = l_Lean_Name_mkStr4(x_99, x_103, x_123, x_138); +lean_inc_ref(x_101); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_139 = l_Lean_Name_mkStr4(x_108, x_105, x_101, x_138); x_140 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_141 = lean_box(2); -lean_inc(x_117); +lean_inc(x_109); x_142 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_117); +lean_ctor_set(x_142, 1, x_109); lean_ctor_set(x_142, 2, x_140); x_143 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -lean_inc(x_119); -x_144 = lean_array_push(x_143, x_119); +lean_inc(x_123); +x_144 = lean_array_push(x_143, x_123); x_145 = lean_array_push(x_144, x_142); lean_inc(x_139); x_146 = lean_alloc_ctor(1, 3, 0); @@ -6946,114 +6942,114 @@ lean_ctor_set(x_146, 0, x_141); lean_ctor_set(x_146, 1, x_139); lean_ctor_set(x_146, 2, x_145); x_147 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_123); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_148 = l_Lean_Name_mkStr4(x_99, x_103, x_123, x_147); -lean_inc_n(x_113, 2); +lean_inc_ref(x_101); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_148 = l_Lean_Name_mkStr4(x_108, x_105, x_101, x_147); +lean_inc_n(x_119, 2); lean_inc(x_148); -lean_inc(x_120); -x_149 = l_Lean_Syntax_node2(x_120, x_148, x_113, x_113); +lean_inc(x_106); +x_149 = l_Lean_Syntax_node2(x_106, x_148, x_119, x_119); x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_104); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_151 = l_Lean_Name_mkStr4(x_99, x_103, x_104, x_150); +lean_inc_ref(x_121); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_151 = l_Lean_Name_mkStr4(x_108, x_105, x_121, x_150); x_152 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__19; x_153 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__21; -lean_inc(x_110); -lean_inc(x_109); -x_154 = l_Lean_addMacroScope(x_109, x_153, x_110); -lean_inc(x_105); +lean_inc(x_99); +lean_inc(x_120); +x_154 = l_Lean_addMacroScope(x_120, x_153, x_99); +lean_inc(x_116); x_155 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_105); -lean_inc(x_130); +lean_ctor_set(x_155, 1, x_116); +lean_inc(x_112); x_156 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_130); -lean_inc(x_120); +lean_ctor_set(x_156, 1, x_112); +lean_inc(x_106); x_157 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_157, 0, x_120); +lean_ctor_set(x_157, 0, x_106); lean_ctor_set(x_157, 1, x_152); lean_ctor_set(x_157, 2, x_154); lean_ctor_set(x_157, 3, x_156); x_158 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__21; -lean_inc_ref(x_104); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_159 = l_Lean_Name_mkStr4(x_99, x_103, x_104, x_158); +lean_inc_ref(x_121); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_159 = l_Lean_Name_mkStr4(x_108, x_105, x_121, x_158); x_160 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__31; -lean_inc_ref(x_104); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_161 = l_Lean_Name_mkStr4(x_99, x_103, x_104, x_160); -lean_inc(x_120); +lean_inc_ref(x_121); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_161 = l_Lean_Name_mkStr4(x_108, x_105, x_121, x_160); +lean_inc(x_106); x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_120); +lean_ctor_set(x_162, 0, x_106); lean_ctor_set(x_162, 1, x_160); x_163 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__32; -lean_inc_ref(x_104); -lean_inc_ref(x_103); -lean_inc_ref(x_99); -x_164 = l_Lean_Name_mkStr4(x_99, x_103, x_104, x_163); -lean_inc(x_117); -lean_inc(x_120); -x_165 = l_Lean_Syntax_node1(x_120, x_117, x_121); +lean_inc_ref(x_121); +lean_inc_ref(x_105); +lean_inc_ref(x_108); +x_164 = l_Lean_Name_mkStr4(x_108, x_105, x_121, x_163); +lean_inc(x_109); +lean_inc(x_106); +x_165 = l_Lean_Syntax_node1(x_106, x_109, x_126); x_166 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__33; -lean_inc(x_120); +lean_inc(x_106); x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_120); +lean_ctor_set(x_167, 0, x_106); lean_ctor_set(x_167, 1, x_166); -lean_inc(x_113); -lean_inc(x_120); -x_168 = l_Lean_Syntax_node4(x_120, x_164, x_165, x_113, x_167, x_122); -lean_inc(x_120); -x_169 = l_Lean_Syntax_node2(x_120, x_161, x_162, x_168); -lean_inc(x_120); -x_170 = l_Lean_Syntax_node3(x_120, x_159, x_106, x_169, x_116); -lean_inc(x_117); -lean_inc(x_120); -x_171 = l_Lean_Syntax_node4(x_120, x_117, x_127, x_114, x_112, x_170); -lean_inc(x_120); -x_172 = l_Lean_Syntax_node2(x_120, x_151, x_157, x_171); -lean_inc_n(x_113, 2); -lean_inc(x_120); -x_173 = l_Lean_Syntax_node2(x_120, x_128, x_113, x_113); -if (lean_obj_tag(x_107) == 0) +lean_inc(x_119); +lean_inc(x_106); +x_168 = l_Lean_Syntax_node4(x_106, x_164, x_165, x_119, x_167, x_128); +lean_inc(x_106); +x_169 = l_Lean_Syntax_node2(x_106, x_161, x_162, x_168); +lean_inc(x_106); +x_170 = l_Lean_Syntax_node3(x_106, x_159, x_124, x_169, x_104); +lean_inc(x_109); +lean_inc(x_106); +x_171 = l_Lean_Syntax_node4(x_106, x_109, x_102, x_129, x_114, x_170); +lean_inc(x_106); +x_172 = l_Lean_Syntax_node2(x_106, x_151, x_157, x_171); +lean_inc_n(x_119, 2); +lean_inc(x_106); +x_173 = l_Lean_Syntax_node2(x_106, x_122, x_119, x_119); +if (lean_obj_tag(x_100) == 0) { lean_object* x_174; x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_4 = x_134; -x_5 = x_99; -x_6 = x_100; +x_4 = x_99; +x_5 = x_134; +x_6 = x_101; x_7 = x_139; -x_8 = x_101; -x_9 = x_102; -x_10 = x_103; -x_11 = x_104; -x_12 = x_105; -x_13 = x_108; -x_14 = x_109; +x_8 = x_103; +x_9 = x_107; +x_10 = x_106; +x_11 = x_105; +x_12 = x_108; +x_13 = x_109; +x_14 = x_110; x_15 = x_172; -x_16 = x_110; -x_17 = x_149; -x_18 = x_136; -x_19 = x_111; -x_20 = x_113; -x_21 = x_115; -x_22 = x_148; -x_23 = x_117; -x_24 = x_118; -x_25 = x_119; -x_26 = x_120; -x_27 = x_173; -x_28 = x_123; -x_29 = x_124; -x_30 = x_126; -x_31 = x_125; -x_32 = x_146; -x_33 = x_129; +x_16 = x_111; +x_17 = x_112; +x_18 = x_149; +x_19 = x_113; +x_20 = x_115; +x_21 = x_116; +x_22 = x_117; +x_23 = x_118; +x_24 = x_119; +x_25 = x_120; +x_26 = x_136; +x_27 = x_121; +x_28 = x_148; +x_29 = x_123; +x_30 = x_125; +x_31 = x_173; +x_32 = x_127; +x_33 = x_146; x_34 = x_137; x_35 = x_130; x_36 = x_174; @@ -7062,40 +7058,40 @@ goto block_89; else { lean_object* x_175; lean_object* x_176; -x_175 = lean_ctor_get(x_107, 0); +x_175 = lean_ctor_get(x_100, 0); lean_inc(x_175); -lean_dec_ref(x_107); +lean_dec_ref(x_100); x_176 = l_Array_mkArray1___redArg(x_175); -x_4 = x_134; -x_5 = x_99; -x_6 = x_100; +x_4 = x_99; +x_5 = x_134; +x_6 = x_101; x_7 = x_139; -x_8 = x_101; -x_9 = x_102; -x_10 = x_103; -x_11 = x_104; -x_12 = x_105; -x_13 = x_108; -x_14 = x_109; +x_8 = x_103; +x_9 = x_107; +x_10 = x_106; +x_11 = x_105; +x_12 = x_108; +x_13 = x_109; +x_14 = x_110; x_15 = x_172; -x_16 = x_110; -x_17 = x_149; -x_18 = x_136; -x_19 = x_111; -x_20 = x_113; -x_21 = x_115; -x_22 = x_148; -x_23 = x_117; -x_24 = x_118; -x_25 = x_119; -x_26 = x_120; -x_27 = x_173; -x_28 = x_123; -x_29 = x_124; -x_30 = x_126; -x_31 = x_125; -x_32 = x_146; -x_33 = x_129; +x_16 = x_111; +x_17 = x_112; +x_18 = x_149; +x_19 = x_113; +x_20 = x_115; +x_21 = x_116; +x_22 = x_117; +x_23 = x_118; +x_24 = x_119; +x_25 = x_120; +x_26 = x_136; +x_27 = x_121; +x_28 = x_148; +x_29 = x_123; +x_30 = x_125; +x_31 = x_173; +x_32 = x_127; +x_33 = x_146; x_34 = x_137; x_35 = x_130; x_36 = x_176; @@ -7119,7 +7115,7 @@ lean_inc(x_199); lean_inc(x_197); lean_inc(x_196); lean_ctor_set(x_193, 5, x_199); -x_200 = l_Lake_DSL_expandOptSimpleBinder(x_191, x_193, x_194); +x_200 = l_Lake_DSL_expandOptSimpleBinder(x_189, x_193, x_194); if (lean_obj_tag(x_200) == 0) { lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; @@ -7132,15 +7128,15 @@ x_203 = 0; x_204 = l_Lean_SourceInfo_fromRef(x_199, x_203); lean_dec(x_199); x_205 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_206 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_205); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_206 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_205); x_207 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_208 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_207); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_208 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_207); x_209 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_210 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_204); @@ -7153,9 +7149,9 @@ lean_inc(x_204); x_212 = l_Lean_Syntax_node1(x_204, x_208, x_211); x_213 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_214 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_215 = l_Lean_Name_mkStr4(x_180, x_183, x_213, x_214); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_215 = l_Lean_Name_mkStr4(x_186, x_185, x_213, x_214); x_216 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_217 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; lean_inc(x_197); @@ -7175,12 +7171,12 @@ lean_inc(x_204); x_222 = l_Lean_Syntax_node2(x_204, x_206, x_212, x_221); x_223 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__53; x_224 = lean_array_push(x_223, x_222); -x_225 = l_Lake_DSL_expandAttrs(x_179); +x_225 = l_Lake_DSL_expandAttrs(x_182); x_226 = l_Array_append___redArg(x_224, x_225); lean_dec_ref(x_225); -x_227 = l_Lean_TSyntax_getId(x_181); -lean_inc(x_181); -x_228 = l_Lake_Name_quoteFrom(x_181, x_227, x_203); +x_227 = l_Lean_TSyntax_getId(x_184); +lean_inc(x_184); +x_228 = l_Lake_Name_quoteFrom(x_184, x_227, x_203); x_229 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; x_230 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; lean_inc(x_204); @@ -7205,15 +7201,15 @@ lean_ctor_set(x_238, 1, x_234); lean_ctor_set(x_238, 2, x_236); lean_ctor_set(x_238, 3, x_237); x_239 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_240 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_239); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_240 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_239); x_241 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_242 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_241); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_242 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_241); x_243 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; lean_inc(x_204); x_244 = lean_alloc_ctor(2, 2, 0); @@ -7226,29 +7222,29 @@ lean_inc(x_197); lean_inc(x_196); x_248 = l_Lean_addMacroScope(x_196, x_247, x_197); x_249 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_250 = l_Lean_Name_mkStr3(x_180, x_183, x_186); +x_250 = l_Lean_Name_mkStr3(x_186, x_185, x_179); x_251 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_251, 0, x_250); x_252 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; +lean_inc_ref(x_179); lean_inc_ref(x_186); -lean_inc_ref(x_180); -x_253 = l_Lean_Name_mkStr3(x_180, x_252, x_186); +x_253 = l_Lean_Name_mkStr3(x_186, x_252, x_179); x_254 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_254, 0, x_253); -lean_inc_ref(x_180); -x_255 = l_Lean_Name_mkStr2(x_180, x_252); +lean_inc_ref(x_186); +x_255 = l_Lean_Name_mkStr2(x_186, x_252); x_256 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_256, 0, x_255); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_257 = l_Lean_Name_mkStr2(x_180, x_183); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_257 = l_Lean_Name_mkStr2(x_186, x_185); x_258 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_258, 0, x_257); -lean_inc_ref(x_180); -x_259 = l_Lean_Name_mkStr1(x_180); +lean_inc_ref(x_186); +x_259 = l_Lean_Name_mkStr1(x_186); x_260 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_260, 0, x_259); x_261 = lean_alloc_ctor(1, 2, 0); @@ -7312,101 +7308,101 @@ lean_inc(x_204); x_282 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_282, 0, x_204); lean_ctor_set(x_282, 1, x_281); -lean_inc(x_189); +lean_inc(x_181); lean_inc_ref(x_282); lean_inc_ref(x_233); -lean_inc(x_181); +lean_inc(x_184); lean_inc_ref(x_211); lean_inc(x_204); -x_283 = l_Lean_Syntax_node8(x_204, x_229, x_211, x_231, x_181, x_233, x_238, x_280, x_282, x_189); +x_283 = l_Lean_Syntax_node8(x_204, x_229, x_211, x_231, x_184, x_233, x_238, x_280, x_282, x_181); x_284 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_285 = l_Lean_Name_mkStr4(x_180, x_183, x_186, x_284); +x_285 = l_Lean_Name_mkStr4(x_186, x_185, x_179, x_284); x_286 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_287 = l_Lean_Name_mkStr4(x_180, x_183, x_186, x_286); -if (lean_obj_tag(x_187) == 0) +x_287 = l_Lean_Name_mkStr4(x_186, x_185, x_179, x_286); +if (lean_obj_tag(x_188) == 0) { lean_object* x_288; x_288 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_99 = x_180; -x_100 = x_226; -x_101 = x_283; -x_102 = x_182; -x_103 = x_183; -x_104 = x_184; -x_105 = x_219; -x_106 = x_269; -x_107 = x_192; -x_108 = x_285; -x_109 = x_196; -x_110 = x_197; -x_111 = x_287; -x_112 = x_228; -x_113 = x_211; -x_114 = x_273; -x_115 = x_282; -x_116 = x_279; -x_117 = x_209; +x_99 = x_197; +x_100 = x_192; +x_101 = x_179; +x_102 = x_181; +x_103 = x_285; +x_104 = x_279; +x_105 = x_185; +x_106 = x_204; +x_107 = x_282; +x_108 = x_186; +x_109 = x_209; +x_110 = x_226; +x_111 = x_274; +x_112 = x_219; +x_113 = x_190; +x_114 = x_228; +x_115 = x_202; +x_116 = x_219; +x_117 = x_283; x_118 = x_210; -x_119 = x_181; -x_120 = x_204; -x_121 = x_201; -x_122 = x_185; -x_123 = x_186; -x_124 = x_274; -x_125 = x_188; -x_126 = x_202; -x_127 = x_189; -x_128 = x_190; -x_129 = x_233; -x_130 = x_219; +x_119 = x_211; +x_120 = x_196; +x_121 = x_180; +x_122 = x_183; +x_123 = x_184; +x_124 = x_269; +x_125 = x_233; +x_126 = x_201; +x_127 = x_287; +x_128 = x_187; +x_129 = x_273; +x_130 = x_191; x_131 = x_288; goto block_177; } else { lean_object* x_289; lean_object* x_290; -x_289 = lean_ctor_get(x_187, 0); +x_289 = lean_ctor_get(x_188, 0); lean_inc(x_289); -lean_dec_ref(x_187); +lean_dec_ref(x_188); x_290 = l_Array_mkArray1___redArg(x_289); -x_99 = x_180; -x_100 = x_226; -x_101 = x_283; -x_102 = x_182; -x_103 = x_183; -x_104 = x_184; -x_105 = x_219; -x_106 = x_269; -x_107 = x_192; -x_108 = x_285; -x_109 = x_196; -x_110 = x_197; -x_111 = x_287; -x_112 = x_228; -x_113 = x_211; -x_114 = x_273; -x_115 = x_282; -x_116 = x_279; -x_117 = x_209; +x_99 = x_197; +x_100 = x_192; +x_101 = x_179; +x_102 = x_181; +x_103 = x_285; +x_104 = x_279; +x_105 = x_185; +x_106 = x_204; +x_107 = x_282; +x_108 = x_186; +x_109 = x_209; +x_110 = x_226; +x_111 = x_274; +x_112 = x_219; +x_113 = x_190; +x_114 = x_228; +x_115 = x_202; +x_116 = x_219; +x_117 = x_283; x_118 = x_210; -x_119 = x_181; -x_120 = x_204; -x_121 = x_201; -x_122 = x_185; -x_123 = x_186; -x_124 = x_274; -x_125 = x_188; -x_126 = x_202; -x_127 = x_189; -x_128 = x_190; -x_129 = x_233; -x_130 = x_219; +x_119 = x_211; +x_120 = x_196; +x_121 = x_180; +x_122 = x_183; +x_123 = x_184; +x_124 = x_269; +x_125 = x_233; +x_126 = x_201; +x_127 = x_287; +x_128 = x_187; +x_129 = x_273; +x_130 = x_191; x_131 = x_290; goto block_177; } @@ -7418,18 +7414,18 @@ lean_dec(x_199); lean_dec(x_197); lean_dec(x_196); lean_dec(x_192); +lean_dec(x_191); lean_dec(x_190); -lean_dec(x_189); lean_dec(x_188); lean_dec(x_187); lean_dec_ref(x_186); -lean_dec(x_185); -lean_dec_ref(x_184); -lean_dec_ref(x_183); +lean_dec_ref(x_185); +lean_dec(x_184); +lean_dec(x_183); lean_dec(x_182); lean_dec(x_181); lean_dec_ref(x_180); -lean_dec(x_179); +lean_dec_ref(x_179); x_291 = !lean_is_exclusive(x_200); if (x_291 == 0) { @@ -7479,7 +7475,7 @@ lean_ctor_set(x_302, 2, x_297); lean_ctor_set(x_302, 3, x_298); lean_ctor_set(x_302, 4, x_299); lean_ctor_set(x_302, 5, x_301); -x_303 = l_Lake_DSL_expandOptSimpleBinder(x_191, x_302, x_194); +x_303 = l_Lake_DSL_expandOptSimpleBinder(x_189, x_302, x_194); if (lean_obj_tag(x_303) == 0) { lean_object* x_304; lean_object* x_305; uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; 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; @@ -7492,15 +7488,15 @@ x_306 = 0; x_307 = l_Lean_SourceInfo_fromRef(x_301, x_306); lean_dec(x_301); x_308 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_309 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_308); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_309 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_308); x_310 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_311 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_310); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_311 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_310); x_312 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_313 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_307); @@ -7513,9 +7509,9 @@ lean_inc(x_307); x_315 = l_Lean_Syntax_node1(x_307, x_311, x_314); x_316 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_317 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_318 = l_Lean_Name_mkStr4(x_180, x_183, x_316, x_317); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_318 = l_Lean_Name_mkStr4(x_186, x_185, x_316, x_317); x_319 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_320 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; lean_inc(x_297); @@ -7535,12 +7531,12 @@ lean_inc(x_307); x_325 = l_Lean_Syntax_node2(x_307, x_309, x_315, x_324); x_326 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__53; x_327 = lean_array_push(x_326, x_325); -x_328 = l_Lake_DSL_expandAttrs(x_179); +x_328 = l_Lake_DSL_expandAttrs(x_182); x_329 = l_Array_append___redArg(x_327, x_328); lean_dec_ref(x_328); -x_330 = l_Lean_TSyntax_getId(x_181); -lean_inc(x_181); -x_331 = l_Lake_Name_quoteFrom(x_181, x_330, x_306); +x_330 = l_Lean_TSyntax_getId(x_184); +lean_inc(x_184); +x_331 = l_Lake_Name_quoteFrom(x_184, x_330, x_306); x_332 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; x_333 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; lean_inc(x_307); @@ -7565,15 +7561,15 @@ lean_ctor_set(x_341, 1, x_337); lean_ctor_set(x_341, 2, x_339); lean_ctor_set(x_341, 3, x_340); x_342 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_343 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_342); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_343 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_342); x_344 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_184); -lean_inc_ref(x_183); lean_inc_ref(x_180); -x_345 = l_Lean_Name_mkStr4(x_180, x_183, x_184, x_344); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_345 = l_Lean_Name_mkStr4(x_186, x_185, x_180, x_344); x_346 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; lean_inc(x_307); x_347 = lean_alloc_ctor(2, 2, 0); @@ -7586,29 +7582,29 @@ lean_inc(x_297); lean_inc(x_296); x_351 = l_Lean_addMacroScope(x_296, x_350, x_297); x_352 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_353 = l_Lean_Name_mkStr3(x_180, x_183, x_186); +x_353 = l_Lean_Name_mkStr3(x_186, x_185, x_179); x_354 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_354, 0, x_353); x_355 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; +lean_inc_ref(x_179); lean_inc_ref(x_186); -lean_inc_ref(x_180); -x_356 = l_Lean_Name_mkStr3(x_180, x_355, x_186); +x_356 = l_Lean_Name_mkStr3(x_186, x_355, x_179); x_357 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_357, 0, x_356); -lean_inc_ref(x_180); -x_358 = l_Lean_Name_mkStr2(x_180, x_355); +lean_inc_ref(x_186); +x_358 = l_Lean_Name_mkStr2(x_186, x_355); x_359 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_359, 0, x_358); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_360 = l_Lean_Name_mkStr2(x_180, x_183); +lean_inc_ref(x_185); +lean_inc_ref(x_186); +x_360 = l_Lean_Name_mkStr2(x_186, x_185); x_361 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_361, 0, x_360); -lean_inc_ref(x_180); -x_362 = l_Lean_Name_mkStr1(x_180); +lean_inc_ref(x_186); +x_362 = l_Lean_Name_mkStr1(x_186); x_363 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_363, 0, x_362); x_364 = lean_alloc_ctor(1, 2, 0); @@ -7672,101 +7668,101 @@ lean_inc(x_307); x_385 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_385, 0, x_307); lean_ctor_set(x_385, 1, x_384); -lean_inc(x_189); +lean_inc(x_181); lean_inc_ref(x_385); lean_inc_ref(x_336); -lean_inc(x_181); +lean_inc(x_184); lean_inc_ref(x_314); lean_inc(x_307); -x_386 = l_Lean_Syntax_node8(x_307, x_332, x_314, x_334, x_181, x_336, x_341, x_383, x_385, x_189); +x_386 = l_Lean_Syntax_node8(x_307, x_332, x_314, x_334, x_184, x_336, x_341, x_383, x_385, x_181); x_387 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_388 = l_Lean_Name_mkStr4(x_180, x_183, x_186, x_387); +x_388 = l_Lean_Name_mkStr4(x_186, x_185, x_179, x_387); x_389 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; +lean_inc_ref(x_179); +lean_inc_ref(x_185); lean_inc_ref(x_186); -lean_inc_ref(x_183); -lean_inc_ref(x_180); -x_390 = l_Lean_Name_mkStr4(x_180, x_183, x_186, x_389); -if (lean_obj_tag(x_187) == 0) +x_390 = l_Lean_Name_mkStr4(x_186, x_185, x_179, x_389); +if (lean_obj_tag(x_188) == 0) { lean_object* x_391; x_391 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_99 = x_180; -x_100 = x_329; -x_101 = x_386; -x_102 = x_182; -x_103 = x_183; -x_104 = x_184; -x_105 = x_322; -x_106 = x_372; -x_107 = x_192; -x_108 = x_388; -x_109 = x_296; -x_110 = x_297; -x_111 = x_390; -x_112 = x_331; -x_113 = x_314; -x_114 = x_376; -x_115 = x_385; -x_116 = x_382; -x_117 = x_312; +x_99 = x_297; +x_100 = x_192; +x_101 = x_179; +x_102 = x_181; +x_103 = x_388; +x_104 = x_382; +x_105 = x_185; +x_106 = x_307; +x_107 = x_385; +x_108 = x_186; +x_109 = x_312; +x_110 = x_329; +x_111 = x_377; +x_112 = x_322; +x_113 = x_190; +x_114 = x_331; +x_115 = x_305; +x_116 = x_322; +x_117 = x_386; x_118 = x_313; -x_119 = x_181; -x_120 = x_307; -x_121 = x_304; -x_122 = x_185; -x_123 = x_186; -x_124 = x_377; -x_125 = x_188; -x_126 = x_305; -x_127 = x_189; -x_128 = x_190; -x_129 = x_336; -x_130 = x_322; +x_119 = x_314; +x_120 = x_296; +x_121 = x_180; +x_122 = x_183; +x_123 = x_184; +x_124 = x_372; +x_125 = x_336; +x_126 = x_304; +x_127 = x_390; +x_128 = x_187; +x_129 = x_376; +x_130 = x_191; x_131 = x_391; goto block_177; } else { lean_object* x_392; lean_object* x_393; -x_392 = lean_ctor_get(x_187, 0); +x_392 = lean_ctor_get(x_188, 0); lean_inc(x_392); -lean_dec_ref(x_187); +lean_dec_ref(x_188); x_393 = l_Array_mkArray1___redArg(x_392); -x_99 = x_180; -x_100 = x_329; -x_101 = x_386; -x_102 = x_182; -x_103 = x_183; -x_104 = x_184; -x_105 = x_322; -x_106 = x_372; -x_107 = x_192; -x_108 = x_388; -x_109 = x_296; -x_110 = x_297; -x_111 = x_390; -x_112 = x_331; -x_113 = x_314; -x_114 = x_376; -x_115 = x_385; -x_116 = x_382; -x_117 = x_312; +x_99 = x_297; +x_100 = x_192; +x_101 = x_179; +x_102 = x_181; +x_103 = x_388; +x_104 = x_382; +x_105 = x_185; +x_106 = x_307; +x_107 = x_385; +x_108 = x_186; +x_109 = x_312; +x_110 = x_329; +x_111 = x_377; +x_112 = x_322; +x_113 = x_190; +x_114 = x_331; +x_115 = x_305; +x_116 = x_322; +x_117 = x_386; x_118 = x_313; -x_119 = x_181; -x_120 = x_307; -x_121 = x_304; -x_122 = x_185; -x_123 = x_186; -x_124 = x_377; -x_125 = x_188; -x_126 = x_305; -x_127 = x_189; -x_128 = x_190; -x_129 = x_336; -x_130 = x_322; +x_119 = x_314; +x_120 = x_296; +x_121 = x_180; +x_122 = x_183; +x_123 = x_184; +x_124 = x_372; +x_125 = x_336; +x_126 = x_304; +x_127 = x_390; +x_128 = x_187; +x_129 = x_376; +x_130 = x_191; x_131 = x_393; goto block_177; } @@ -7778,18 +7774,18 @@ lean_dec(x_301); lean_dec(x_297); lean_dec(x_296); lean_dec(x_192); +lean_dec(x_191); lean_dec(x_190); -lean_dec(x_189); lean_dec(x_188); lean_dec(x_187); lean_dec_ref(x_186); -lean_dec(x_185); -lean_dec_ref(x_184); -lean_dec_ref(x_183); +lean_dec_ref(x_185); +lean_dec(x_184); +lean_dec(x_183); lean_dec(x_182); lean_dec(x_181); lean_dec_ref(x_180); -lean_dec(x_179); +lean_dec_ref(x_179); x_394 = lean_ctor_get(x_303, 0); lean_inc(x_394); x_395 = lean_ctor_get(x_303, 1); @@ -7985,19 +7981,19 @@ lean_object* x_446; lean_dec(x_400); x_446 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_446, 0, x_441); -x_179 = x_401; -x_180 = x_408; -x_181 = x_403; -x_182 = x_411; -x_183 = x_409; -x_184 = x_410; -x_185 = x_435; -x_186 = x_416; -x_187 = x_402; -x_188 = x_417; -x_189 = x_434; -x_190 = x_422; -x_191 = x_404; +x_179 = x_416; +x_180 = x_410; +x_181 = x_434; +x_182 = x_402; +x_183 = x_422; +x_184 = x_403; +x_185 = x_409; +x_186 = x_408; +x_187 = x_435; +x_188 = x_401; +x_189 = x_404; +x_190 = x_411; +x_191 = x_417; x_192 = x_446; x_193 = x_405; x_194 = x_406; @@ -8011,19 +8007,19 @@ lean_object* x_447; lean_dec(x_436); lean_dec(x_400); x_447 = lean_box(0); -x_179 = x_401; -x_180 = x_408; -x_181 = x_403; -x_182 = x_411; -x_183 = x_409; -x_184 = x_410; -x_185 = x_435; -x_186 = x_416; -x_187 = x_402; -x_188 = x_417; -x_189 = x_434; -x_190 = x_422; -x_191 = x_404; +x_179 = x_416; +x_180 = x_410; +x_181 = x_434; +x_182 = x_402; +x_183 = x_422; +x_184 = x_403; +x_185 = x_409; +x_186 = x_408; +x_187 = x_435; +x_188 = x_401; +x_189 = x_404; +x_190 = x_411; +x_191 = x_417; x_192 = x_447; x_193 = x_405; x_194 = x_406; @@ -8120,8 +8116,8 @@ x_470 = l_Lean_Syntax_getArg(x_465, x_94); lean_dec(x_465); x_471 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_471, 0, x_470); -x_401 = x_450; -x_402 = x_451; +x_401 = x_451; +x_402 = x_450; x_403 = x_460; x_404 = x_471; x_405 = x_2; @@ -8134,8 +8130,8 @@ else lean_object* x_472; lean_dec(x_465); x_472 = lean_box(0); -x_401 = x_450; -x_402 = x_451; +x_401 = x_451; +x_402 = x_450; x_403 = x_460; x_404 = x_472; x_405 = x_2; @@ -8187,142 +8183,142 @@ goto block_473; block_89: { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_inc_ref(x_24); -x_37 = l_Array_append___redArg(x_24, x_36); +lean_inc_ref(x_23); +x_37 = l_Array_append___redArg(x_23, x_36); lean_dec_ref(x_36); -lean_inc(x_23); -lean_inc(x_26); +lean_inc(x_13); +lean_inc(x_10); x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_26); -lean_ctor_set(x_38, 1, x_23); +lean_ctor_set(x_38, 0, x_10); +lean_ctor_set(x_38, 1, x_13); lean_ctor_set(x_38, 2, x_37); -lean_inc(x_27); -lean_inc(x_21); lean_inc(x_31); -lean_inc(x_26); -x_39 = l_Lean_Syntax_node4(x_26, x_31, x_21, x_15, x_27, x_38); -lean_inc(x_26); -x_40 = l_Lean_Syntax_node4(x_26, x_18, x_34, x_32, x_17, x_39); -lean_inc(x_13); -lean_inc(x_26); -x_41 = l_Lean_Syntax_node2(x_26, x_13, x_4, x_40); +lean_inc(x_9); +lean_inc(x_35); +lean_inc(x_10); +x_39 = l_Lean_Syntax_node4(x_10, x_35, x_9, x_15, x_31, x_38); +lean_inc(x_10); +x_40 = l_Lean_Syntax_node4(x_10, x_26, x_34, x_33, x_18, x_39); +lean_inc(x_8); +lean_inc(x_10); +x_41 = l_Lean_Syntax_node2(x_10, x_8, x_5, x_40); x_42 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; +lean_inc_ref(x_27); lean_inc_ref(x_11); -lean_inc_ref(x_10); -lean_inc_ref(x_5); -x_43 = l_Lean_Name_mkStr4(x_5, x_10, x_11, x_42); +lean_inc_ref(x_12); +x_43 = l_Lean_Name_mkStr4(x_12, x_11, x_27, x_42); x_44 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_26); +lean_inc(x_10); x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_26); +lean_ctor_set(x_45, 0, x_10); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Syntax_SepArray_ofElems(x_29, x_6); -lean_dec_ref(x_6); -x_47 = l_Array_append___redArg(x_24, x_46); +x_46 = l_Lean_Syntax_SepArray_ofElems(x_16, x_14); +lean_dec_ref(x_14); +x_47 = l_Array_append___redArg(x_23, x_46); lean_dec_ref(x_46); -lean_inc(x_23); -lean_inc(x_26); +lean_inc(x_13); +lean_inc(x_10); x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_26); -lean_ctor_set(x_48, 1, x_23); +lean_ctor_set(x_48, 0, x_10); +lean_ctor_set(x_48, 1, x_13); lean_ctor_set(x_48, 2, x_47); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_26); +lean_inc(x_10); x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_26); +lean_ctor_set(x_50, 0, x_10); lean_ctor_set(x_50, 1, x_49); -lean_inc(x_26); -x_51 = l_Lean_Syntax_node3(x_26, x_43, x_45, x_48, x_50); -lean_inc(x_23); -lean_inc(x_26); -x_52 = l_Lean_Syntax_node1(x_26, x_23, x_51); -lean_inc_n(x_20, 6); -lean_inc(x_26); -x_53 = l_Lean_Syntax_node7(x_26, x_19, x_20, x_52, x_20, x_20, x_20, x_20, x_20); +lean_inc(x_10); +x_51 = l_Lean_Syntax_node3(x_10, x_43, x_45, x_48, x_50); +lean_inc(x_13); +lean_inc(x_10); +x_52 = l_Lean_Syntax_node1(x_10, x_13, x_51); +lean_inc_n(x_24, 6); +lean_inc(x_10); +x_53 = l_Lean_Syntax_node7(x_10, x_32, x_24, x_52, x_24, x_24, x_24, x_24, x_24); x_54 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; -lean_inc_ref(x_10); -lean_inc_ref(x_5); -x_55 = l_Lean_Name_mkStr4(x_5, x_10, x_28, x_54); +lean_inc_ref(x_11); +lean_inc_ref(x_12); +x_55 = l_Lean_Name_mkStr4(x_12, x_11, x_6, x_54); x_56 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_26); +lean_inc(x_10); x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_26); +lean_ctor_set(x_57, 0, x_10); lean_ctor_set(x_57, 1, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_59 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_16); -lean_inc(x_14); -x_60 = l_Lean_addMacroScope(x_14, x_59, x_16); -lean_inc(x_35); -lean_inc(x_26); +lean_inc(x_4); +lean_inc(x_25); +x_60 = l_Lean_addMacroScope(x_25, x_59, x_4); +lean_inc(x_17); +lean_inc(x_10); x_61 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_61, 0, x_26); +lean_ctor_set(x_61, 0, x_10); lean_ctor_set(x_61, 1, x_58); lean_ctor_set(x_61, 2, x_60); -lean_ctor_set(x_61, 3, x_35); -lean_inc(x_20); -lean_inc(x_26); -x_62 = l_Lean_Syntax_node2(x_26, x_7, x_61, x_20); +lean_ctor_set(x_61, 3, x_17); +lean_inc(x_24); +lean_inc(x_10); +x_62 = l_Lean_Syntax_node2(x_10, x_7, x_61, x_24); x_63 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_64 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_16); -lean_inc(x_14); -x_65 = l_Lean_addMacroScope(x_14, x_64, x_16); +lean_inc(x_4); +lean_inc(x_25); +x_65 = l_Lean_addMacroScope(x_25, x_64, x_4); x_66 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__8; x_67 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_12); +lean_ctor_set(x_67, 1, x_21); x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__9; -lean_inc(x_35); +lean_inc(x_17); x_69 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_35); +lean_ctor_set(x_69, 1, x_17); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_67); lean_ctor_set(x_70, 1, x_69); -lean_inc(x_26); +lean_inc(x_10); x_71 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_71, 0, x_26); +lean_ctor_set(x_71, 0, x_10); lean_ctor_set(x_71, 1, x_63); lean_ctor_set(x_71, 2, x_65); lean_ctor_set(x_71, 3, x_70); -lean_inc(x_26); -x_72 = l_Lean_Syntax_node2(x_26, x_9, x_33, x_71); -lean_inc(x_23); -lean_inc(x_26); -x_73 = l_Lean_Syntax_node1(x_26, x_23, x_72); -lean_inc(x_20); -lean_inc(x_26); -x_74 = l_Lean_Syntax_node2(x_26, x_22, x_20, x_73); +lean_inc(x_10); +x_72 = l_Lean_Syntax_node2(x_10, x_19, x_30, x_71); +lean_inc(x_13); +lean_inc(x_10); +x_73 = l_Lean_Syntax_node1(x_10, x_13, x_72); +lean_inc(x_24); +lean_inc(x_10); +x_74 = l_Lean_Syntax_node2(x_10, x_28, x_24, x_73); x_75 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_76 = l_Lean_Name_mkStr4(x_5, x_10, x_11, x_75); +x_76 = l_Lean_Name_mkStr4(x_12, x_11, x_27, x_75); x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_26); +lean_inc(x_10); x_78 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_78, 0, x_26); +lean_ctor_set(x_78, 0, x_10); lean_ctor_set(x_78, 1, x_77); x_79 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_80 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_81 = l_Lean_addMacroScope(x_14, x_80, x_16); -lean_inc(x_26); +x_81 = l_Lean_addMacroScope(x_25, x_80, x_4); +lean_inc(x_10); x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_26); +lean_ctor_set(x_82, 0, x_10); lean_ctor_set(x_82, 1, x_79); lean_ctor_set(x_82, 2, x_81); -lean_ctor_set(x_82, 3, x_35); -lean_inc(x_26); -x_83 = l_Lean_Syntax_node3(x_26, x_76, x_25, x_78, x_82); -lean_inc(x_20); -lean_inc(x_26); -x_84 = l_Lean_Syntax_node4(x_26, x_31, x_21, x_83, x_27, x_20); -lean_inc(x_26); -x_85 = l_Lean_Syntax_node5(x_26, x_55, x_57, x_62, x_74, x_84, x_20); -lean_inc(x_26); -x_86 = l_Lean_Syntax_node2(x_26, x_13, x_53, x_85); -x_87 = l_Lean_Syntax_node3(x_26, x_23, x_8, x_41, x_86); +lean_ctor_set(x_82, 3, x_17); +lean_inc(x_10); +x_83 = l_Lean_Syntax_node3(x_10, x_76, x_29, x_78, x_82); +lean_inc(x_24); +lean_inc(x_10); +x_84 = l_Lean_Syntax_node4(x_10, x_35, x_9, x_83, x_31, x_24); +lean_inc(x_10); +x_85 = l_Lean_Syntax_node5(x_10, x_55, x_57, x_62, x_74, x_84, x_24); +lean_inc(x_10); +x_86 = l_Lean_Syntax_node2(x_10, x_8, x_53, x_85); +x_87 = l_Lean_Syntax_node3(x_10, x_13, x_22, x_41, x_86); x_88 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_30); +lean_ctor_set(x_88, 1, x_20); return x_88; } } @@ -8793,22 +8789,22 @@ lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; x_281 = lean_ctor_get(x_280, 0); lean_inc(x_281); lean_dec_ref(x_280); -lean_inc_ref(x_268); -lean_inc(x_257); -lean_inc(x_265); +lean_inc_ref(x_261); +lean_inc(x_271); +lean_inc(x_269); x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_265); -lean_ctor_set(x_282, 1, x_257); -lean_ctor_set(x_282, 2, x_268); +lean_ctor_set(x_282, 0, x_269); +lean_ctor_set(x_282, 1, x_271); +lean_ctor_set(x_282, 2, x_261); lean_inc_ref(x_282); -lean_inc(x_265); -x_283 = l_Lean_Syntax_node1(x_265, x_269, x_282); -lean_inc(x_265); -x_284 = l_Lean_Syntax_node2(x_265, x_259, x_263, x_282); -x_285 = l_Lean_Syntax_node2(x_265, x_270, x_283, x_284); +lean_inc(x_269); +x_283 = l_Lean_Syntax_node1(x_269, x_265, x_282); +lean_inc(x_269); +x_284 = l_Lean_Syntax_node2(x_269, x_257, x_263, x_282); +x_285 = l_Lean_Syntax_node2(x_269, x_258, x_283, x_284); x_286 = lean_unsigned_to_nat(2u); x_287 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_288 = lean_array_push(x_287, x_266); +x_288 = lean_array_push(x_287, x_260); x_289 = lean_array_push(x_288, x_285); x_290 = l_Lake_DSL_expandAttrs(x_7); x_291 = l_Array_append___redArg(x_289, x_290); @@ -8825,21 +8821,21 @@ lean_object* x_295; x_295 = lean_ctor_get(x_294, 0); lean_inc(x_295); lean_dec_ref(x_294); -x_133 = x_275; -x_134 = x_257; -x_135 = x_258; +x_133 = x_259; +x_134 = x_261; +x_135 = x_262; x_136 = x_291; -x_137 = x_286; -x_138 = x_279; -x_139 = x_260; -x_140 = x_261; -x_141 = x_262; -x_142 = x_264; -x_143 = x_281; -x_144 = x_267; -x_145 = x_268; -x_146 = x_271; -x_147 = x_292; +x_137 = x_264; +x_138 = x_281; +x_139 = x_292; +x_140 = x_266; +x_141 = x_268; +x_142 = x_267; +x_143 = x_286; +x_144 = x_275; +x_145 = x_270; +x_146 = x_279; +x_147 = x_271; x_148 = x_295; x_149 = lean_box(0); goto block_230; @@ -8852,15 +8848,15 @@ lean_dec_ref(x_291); lean_dec(x_281); lean_dec(x_279); lean_dec(x_275); -lean_dec_ref(x_271); -lean_dec_ref(x_268); -lean_dec(x_267); +lean_dec(x_271); +lean_dec_ref(x_270); +lean_dec(x_268); +lean_dec_ref(x_267); +lean_dec_ref(x_266); lean_dec(x_264); -lean_dec_ref(x_262); -lean_dec(x_261); -lean_dec(x_260); -lean_dec_ref(x_258); -lean_dec(x_257); +lean_dec(x_262); +lean_dec_ref(x_261); +lean_dec(x_259); lean_dec(x_6); x_296 = !lean_is_exclusive(x_294); if (x_296 == 0) @@ -8889,21 +8885,21 @@ lean_dec_ref(x_10); x_299 = lean_ctor_get(x_255, 0); lean_inc(x_299); lean_dec_ref(x_255); -x_133 = x_275; -x_134 = x_257; -x_135 = x_258; +x_133 = x_259; +x_134 = x_261; +x_135 = x_262; x_136 = x_291; -x_137 = x_286; -x_138 = x_279; -x_139 = x_260; -x_140 = x_261; -x_141 = x_262; -x_142 = x_264; -x_143 = x_281; -x_144 = x_267; -x_145 = x_268; -x_146 = x_271; -x_147 = x_292; +x_137 = x_264; +x_138 = x_281; +x_139 = x_292; +x_140 = x_266; +x_141 = x_268; +x_142 = x_267; +x_143 = x_286; +x_144 = x_275; +x_145 = x_270; +x_146 = x_279; +x_147 = x_271; x_148 = x_299; x_149 = lean_box(0); goto block_230; @@ -8915,20 +8911,20 @@ uint8_t x_300; lean_dec(x_279); lean_dec(x_277); lean_dec(x_275); -lean_dec_ref(x_271); -lean_dec(x_270); +lean_dec(x_271); +lean_dec_ref(x_270); lean_dec(x_269); -lean_dec_ref(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_dec(x_268); +lean_dec_ref(x_267); +lean_dec_ref(x_266); lean_dec(x_265); lean_dec(x_264); lean_dec(x_263); -lean_dec_ref(x_262); -lean_dec(x_261); +lean_dec(x_262); +lean_dec_ref(x_261); lean_dec(x_260); lean_dec(x_259); -lean_dec_ref(x_258); +lean_dec(x_258); lean_dec(x_257); lean_dec_ref(x_231); lean_dec(x_11); @@ -8958,20 +8954,20 @@ else uint8_t x_303; lean_dec(x_277); lean_dec(x_275); -lean_dec_ref(x_271); -lean_dec(x_270); +lean_dec(x_271); +lean_dec_ref(x_270); lean_dec(x_269); -lean_dec_ref(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_dec(x_268); +lean_dec_ref(x_267); +lean_dec_ref(x_266); lean_dec(x_265); lean_dec(x_264); lean_dec(x_263); -lean_dec_ref(x_262); -lean_dec(x_261); +lean_dec(x_262); +lean_dec_ref(x_261); lean_dec(x_260); lean_dec(x_259); -lean_dec_ref(x_258); +lean_dec(x_258); lean_dec(x_257); lean_dec_ref(x_231); lean_dec(x_11); @@ -9000,20 +8996,20 @@ else { uint8_t x_306; lean_dec(x_275); -lean_dec_ref(x_271); -lean_dec(x_270); +lean_dec(x_271); +lean_dec_ref(x_270); lean_dec(x_269); -lean_dec_ref(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_dec(x_268); +lean_dec_ref(x_267); +lean_dec_ref(x_266); lean_dec(x_265); lean_dec(x_264); lean_dec(x_263); -lean_dec_ref(x_262); -lean_dec(x_261); +lean_dec(x_262); +lean_dec_ref(x_261); lean_dec(x_260); lean_dec(x_259); -lean_dec_ref(x_258); +lean_dec(x_258); lean_dec(x_257); lean_dec_ref(x_231); lean_dec(x_11); @@ -9040,20 +9036,20 @@ return x_308; } else { -lean_dec_ref(x_271); -lean_dec(x_270); +lean_dec(x_271); +lean_dec_ref(x_270); lean_dec(x_269); -lean_dec_ref(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_dec(x_268); +lean_dec_ref(x_267); +lean_dec_ref(x_266); lean_dec(x_265); lean_dec(x_264); lean_dec(x_263); -lean_dec_ref(x_262); -lean_dec(x_261); +lean_dec(x_262); +lean_dec_ref(x_261); lean_dec(x_260); lean_dec(x_259); -lean_dec_ref(x_258); +lean_dec(x_258); lean_dec(x_257); lean_dec_ref(x_231); lean_dec(x_11); @@ -9086,28 +9082,28 @@ x_325 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_326 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_327 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_328 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_310); +lean_inc(x_315); x_329 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_329, 0, x_310); +lean_ctor_set(x_329, 0, x_315); lean_ctor_set(x_329, 1, x_327); lean_ctor_set(x_329, 2, x_328); lean_inc_ref(x_329); -lean_inc(x_310); -x_330 = l_Lean_Syntax_node1(x_310, x_326, x_329); +lean_inc(x_315); +x_330 = l_Lean_Syntax_node1(x_315, x_326, x_329); x_331 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_332 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_333 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_334 = l_Lean_addMacroScope(x_317, x_333, x_314); -lean_inc(x_316); -lean_inc(x_310); +x_334 = l_Lean_addMacroScope(x_317, x_333, x_313); +lean_inc(x_314); +lean_inc(x_315); x_335 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_335, 0, x_310); +lean_ctor_set(x_335, 0, x_315); lean_ctor_set(x_335, 1, x_332); lean_ctor_set(x_335, 2, x_334); -lean_ctor_set(x_335, 3, x_316); -lean_inc(x_310); -x_336 = l_Lean_Syntax_node2(x_310, x_331, x_335, x_329); -x_337 = l_Lean_Syntax_node2(x_310, x_325, x_330, x_336); +lean_ctor_set(x_335, 3, x_314); +lean_inc(x_315); +x_336 = l_Lean_Syntax_node2(x_315, x_331, x_335, x_329); +x_337 = l_Lean_Syntax_node2(x_315, x_325, x_330, x_336); if (lean_obj_tag(x_255) == 0) { lean_object* x_338; lean_object* x_339; @@ -9119,21 +9115,21 @@ x_339 = lean_apply_3(x_338, x_10, x_11, lean_box(0)); if (lean_obj_tag(x_339) == 0) { lean_dec_ref(x_339); -x_257 = x_327; -x_258 = x_322; -x_259 = x_331; -x_260 = x_315; -x_261 = x_316; -x_262 = x_323; +x_257 = x_331; +x_258 = x_325; +x_259 = x_312; +x_260 = x_337; +x_261 = x_328; +x_262 = x_316; x_263 = x_311; -x_264 = x_312; -x_265 = x_320; -x_266 = x_337; -x_267 = x_313; -x_268 = x_328; -x_269 = x_326; -x_270 = x_325; -x_271 = x_324; +x_264 = x_310; +x_265 = x_326; +x_266 = x_323; +x_267 = x_324; +x_268 = x_314; +x_269 = x_320; +x_270 = x_322; +x_271 = x_327; x_272 = lean_box(0); goto block_309; } @@ -9143,10 +9139,10 @@ uint8_t x_340; lean_dec(x_337); lean_dec(x_320); lean_dec(x_316); -lean_dec(x_315); -lean_dec(x_313); +lean_dec(x_314); lean_dec(x_312); lean_dec(x_311); +lean_dec(x_310); lean_dec_ref(x_231); lean_dec(x_11); lean_dec_ref(x_10); @@ -9173,21 +9169,21 @@ return x_342; } else { -x_257 = x_327; -x_258 = x_322; -x_259 = x_331; -x_260 = x_315; -x_261 = x_316; -x_262 = x_323; +x_257 = x_331; +x_258 = x_325; +x_259 = x_312; +x_260 = x_337; +x_261 = x_328; +x_262 = x_316; x_263 = x_311; -x_264 = x_312; -x_265 = x_320; -x_266 = x_337; -x_267 = x_313; -x_268 = x_328; -x_269 = x_326; -x_270 = x_325; -x_271 = x_324; +x_264 = x_310; +x_265 = x_326; +x_266 = x_323; +x_267 = x_324; +x_268 = x_314; +x_269 = x_320; +x_270 = x_322; +x_271 = x_327; x_272 = lean_box(0); goto block_309; } @@ -9338,13 +9334,13 @@ lean_object* x_375; x_375 = lean_ctor_get(x_374, 0); lean_inc(x_375); lean_dec_ref(x_374); -x_310 = x_370; +x_310 = x_354; x_311 = x_368; -x_312 = x_359; -x_313 = x_361; -x_314 = x_372; -x_315 = x_354; -x_316 = x_358; +x_312 = x_361; +x_313 = x_372; +x_314 = x_358; +x_315 = x_370; +x_316 = x_359; x_317 = x_375; x_318 = lean_box(0); goto block_349; @@ -9390,13 +9386,13 @@ lean_inc(x_379); lean_dec_ref(x_371); x_380 = lean_ctor_get(x_255, 0); lean_inc(x_380); -x_310 = x_370; +x_310 = x_354; x_311 = x_368; -x_312 = x_359; -x_313 = x_361; -x_314 = x_379; -x_315 = x_354; -x_316 = x_358; +x_312 = x_361; +x_313 = x_379; +x_314 = x_358; +x_315 = x_370; +x_316 = x_359; x_317 = x_380; x_318 = lean_box(0); goto block_349; @@ -9749,22 +9745,22 @@ lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; x_449 = lean_ctor_get(x_448, 0); lean_inc(x_449); lean_dec_ref(x_448); -lean_inc_ref(x_436); -lean_inc(x_425); -lean_inc(x_433); +lean_inc_ref(x_429); +lean_inc(x_439); +lean_inc(x_437); x_450 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_450, 0, x_433); -lean_ctor_set(x_450, 1, x_425); -lean_ctor_set(x_450, 2, x_436); +lean_ctor_set(x_450, 0, x_437); +lean_ctor_set(x_450, 1, x_439); +lean_ctor_set(x_450, 2, x_429); lean_inc_ref(x_450); -lean_inc(x_433); -x_451 = l_Lean_Syntax_node1(x_433, x_437, x_450); -lean_inc(x_433); -x_452 = l_Lean_Syntax_node2(x_433, x_427, x_431, x_450); -x_453 = l_Lean_Syntax_node2(x_433, x_438, x_451, x_452); +lean_inc(x_437); +x_451 = l_Lean_Syntax_node1(x_437, x_433, x_450); +lean_inc(x_437); +x_452 = l_Lean_Syntax_node2(x_437, x_425, x_431, x_450); +x_453 = l_Lean_Syntax_node2(x_437, x_426, x_451, x_452); x_454 = lean_unsigned_to_nat(2u); x_455 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_456 = lean_array_push(x_455, x_434); +x_456 = lean_array_push(x_455, x_428); x_457 = lean_array_push(x_456, x_453); x_458 = l_Lake_DSL_expandAttrs(x_7); x_459 = l_Array_append___redArg(x_457, x_458); @@ -9781,21 +9777,21 @@ lean_object* x_463; x_463 = lean_ctor_get(x_462, 0); lean_inc(x_463); lean_dec_ref(x_462); -x_133 = x_443; -x_134 = x_425; -x_135 = x_426; +x_133 = x_427; +x_134 = x_429; +x_135 = x_430; x_136 = x_459; -x_137 = x_454; -x_138 = x_447; -x_139 = x_428; -x_140 = x_429; -x_141 = x_430; -x_142 = x_432; -x_143 = x_449; -x_144 = x_435; -x_145 = x_436; -x_146 = x_439; -x_147 = x_460; +x_137 = x_432; +x_138 = x_449; +x_139 = x_460; +x_140 = x_434; +x_141 = x_436; +x_142 = x_435; +x_143 = x_454; +x_144 = x_443; +x_145 = x_438; +x_146 = x_447; +x_147 = x_439; x_148 = x_463; x_149 = lean_box(0); goto block_230; @@ -9808,15 +9804,15 @@ lean_dec_ref(x_459); lean_dec(x_449); lean_dec(x_447); lean_dec(x_443); -lean_dec_ref(x_439); -lean_dec_ref(x_436); -lean_dec(x_435); +lean_dec(x_439); +lean_dec_ref(x_438); +lean_dec(x_436); +lean_dec_ref(x_435); +lean_dec_ref(x_434); lean_dec(x_432); -lean_dec_ref(x_430); -lean_dec(x_429); -lean_dec(x_428); -lean_dec_ref(x_426); -lean_dec(x_425); +lean_dec(x_430); +lean_dec_ref(x_429); +lean_dec(x_427); lean_dec(x_6); x_464 = lean_ctor_get(x_462, 0); lean_inc(x_464); @@ -9846,21 +9842,21 @@ lean_dec_ref(x_10); x_467 = lean_ctor_get(x_423, 0); lean_inc(x_467); lean_dec_ref(x_423); -x_133 = x_443; -x_134 = x_425; -x_135 = x_426; +x_133 = x_427; +x_134 = x_429; +x_135 = x_430; x_136 = x_459; -x_137 = x_454; -x_138 = x_447; -x_139 = x_428; -x_140 = x_429; -x_141 = x_430; -x_142 = x_432; -x_143 = x_449; -x_144 = x_435; -x_145 = x_436; -x_146 = x_439; -x_147 = x_460; +x_137 = x_432; +x_138 = x_449; +x_139 = x_460; +x_140 = x_434; +x_141 = x_436; +x_142 = x_435; +x_143 = x_454; +x_144 = x_443; +x_145 = x_438; +x_146 = x_447; +x_147 = x_439; x_148 = x_467; x_149 = lean_box(0); goto block_230; @@ -9872,20 +9868,20 @@ lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_dec(x_447); lean_dec(x_445); lean_dec(x_443); -lean_dec_ref(x_439); -lean_dec(x_438); +lean_dec(x_439); +lean_dec_ref(x_438); lean_dec(x_437); -lean_dec_ref(x_436); -lean_dec(x_435); -lean_dec(x_434); +lean_dec(x_436); +lean_dec_ref(x_435); +lean_dec_ref(x_434); lean_dec(x_433); lean_dec(x_432); lean_dec(x_431); -lean_dec_ref(x_430); -lean_dec(x_429); +lean_dec(x_430); +lean_dec_ref(x_429); lean_dec(x_428); lean_dec(x_427); -lean_dec_ref(x_426); +lean_dec(x_426); lean_dec(x_425); lean_dec_ref(x_231); lean_dec(x_11); @@ -9916,20 +9912,20 @@ else lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_dec(x_445); lean_dec(x_443); -lean_dec_ref(x_439); -lean_dec(x_438); +lean_dec(x_439); +lean_dec_ref(x_438); lean_dec(x_437); -lean_dec_ref(x_436); -lean_dec(x_435); -lean_dec(x_434); +lean_dec(x_436); +lean_dec_ref(x_435); +lean_dec_ref(x_434); lean_dec(x_433); lean_dec(x_432); lean_dec(x_431); -lean_dec_ref(x_430); -lean_dec(x_429); +lean_dec(x_430); +lean_dec_ref(x_429); lean_dec(x_428); lean_dec(x_427); -lean_dec_ref(x_426); +lean_dec(x_426); lean_dec(x_425); lean_dec_ref(x_231); lean_dec(x_11); @@ -9959,20 +9955,20 @@ else { lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_dec(x_443); -lean_dec_ref(x_439); -lean_dec(x_438); +lean_dec(x_439); +lean_dec_ref(x_438); lean_dec(x_437); -lean_dec_ref(x_436); -lean_dec(x_435); -lean_dec(x_434); +lean_dec(x_436); +lean_dec_ref(x_435); +lean_dec_ref(x_434); lean_dec(x_433); lean_dec(x_432); lean_dec(x_431); -lean_dec_ref(x_430); -lean_dec(x_429); +lean_dec(x_430); +lean_dec_ref(x_429); lean_dec(x_428); lean_dec(x_427); -lean_dec_ref(x_426); +lean_dec(x_426); lean_dec(x_425); lean_dec_ref(x_231); lean_dec(x_11); @@ -10000,20 +9996,20 @@ return x_476; } else { -lean_dec_ref(x_439); -lean_dec(x_438); +lean_dec(x_439); +lean_dec_ref(x_438); lean_dec(x_437); -lean_dec_ref(x_436); -lean_dec(x_435); -lean_dec(x_434); +lean_dec(x_436); +lean_dec_ref(x_435); +lean_dec_ref(x_434); lean_dec(x_433); lean_dec(x_432); lean_dec(x_431); -lean_dec_ref(x_430); -lean_dec(x_429); +lean_dec(x_430); +lean_dec_ref(x_429); lean_dec(x_428); lean_dec(x_427); -lean_dec_ref(x_426); +lean_dec(x_426); lean_dec(x_425); lean_dec_ref(x_231); lean_dec(x_11); @@ -10046,28 +10042,28 @@ x_493 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_494 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_495 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_496 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_478); +lean_inc(x_483); x_497 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_497, 0, x_478); +lean_ctor_set(x_497, 0, x_483); lean_ctor_set(x_497, 1, x_495); lean_ctor_set(x_497, 2, x_496); lean_inc_ref(x_497); -lean_inc(x_478); -x_498 = l_Lean_Syntax_node1(x_478, x_494, x_497); +lean_inc(x_483); +x_498 = l_Lean_Syntax_node1(x_483, x_494, x_497); x_499 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_500 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_501 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_502 = l_Lean_addMacroScope(x_485, x_501, x_482); -lean_inc(x_484); -lean_inc(x_478); +x_502 = l_Lean_addMacroScope(x_485, x_501, x_481); +lean_inc(x_482); +lean_inc(x_483); x_503 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_503, 0, x_478); +lean_ctor_set(x_503, 0, x_483); lean_ctor_set(x_503, 1, x_500); lean_ctor_set(x_503, 2, x_502); -lean_ctor_set(x_503, 3, x_484); -lean_inc(x_478); -x_504 = l_Lean_Syntax_node2(x_478, x_499, x_503, x_497); -x_505 = l_Lean_Syntax_node2(x_478, x_493, x_498, x_504); +lean_ctor_set(x_503, 3, x_482); +lean_inc(x_483); +x_504 = l_Lean_Syntax_node2(x_483, x_499, x_503, x_497); +x_505 = l_Lean_Syntax_node2(x_483, x_493, x_498, x_504); if (lean_obj_tag(x_423) == 0) { lean_object* x_506; lean_object* x_507; @@ -10079,21 +10075,21 @@ x_507 = lean_apply_3(x_506, x_10, x_11, lean_box(0)); if (lean_obj_tag(x_507) == 0) { lean_dec_ref(x_507); -x_425 = x_495; -x_426 = x_490; -x_427 = x_499; -x_428 = x_483; -x_429 = x_484; -x_430 = x_491; +x_425 = x_499; +x_426 = x_493; +x_427 = x_480; +x_428 = x_505; +x_429 = x_496; +x_430 = x_484; x_431 = x_479; -x_432 = x_480; -x_433 = x_488; -x_434 = x_505; -x_435 = x_481; -x_436 = x_496; -x_437 = x_494; -x_438 = x_493; -x_439 = x_492; +x_432 = x_478; +x_433 = x_494; +x_434 = x_491; +x_435 = x_492; +x_436 = x_482; +x_437 = x_488; +x_438 = x_490; +x_439 = x_495; x_440 = lean_box(0); goto block_477; } @@ -10103,10 +10099,10 @@ lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_dec(x_505); lean_dec(x_488); lean_dec(x_484); -lean_dec(x_483); -lean_dec(x_481); +lean_dec(x_482); lean_dec(x_480); lean_dec(x_479); +lean_dec(x_478); lean_dec_ref(x_231); lean_dec(x_11); lean_dec_ref(x_10); @@ -10134,21 +10130,21 @@ return x_510; } else { -x_425 = x_495; -x_426 = x_490; -x_427 = x_499; -x_428 = x_483; -x_429 = x_484; -x_430 = x_491; +x_425 = x_499; +x_426 = x_493; +x_427 = x_480; +x_428 = x_505; +x_429 = x_496; +x_430 = x_484; x_431 = x_479; -x_432 = x_480; -x_433 = x_488; -x_434 = x_505; -x_435 = x_481; -x_436 = x_496; -x_437 = x_494; -x_438 = x_493; -x_439 = x_492; +x_432 = x_478; +x_433 = x_494; +x_434 = x_491; +x_435 = x_492; +x_436 = x_482; +x_437 = x_488; +x_438 = x_490; +x_439 = x_495; x_440 = lean_box(0); goto block_477; } @@ -10301,13 +10297,13 @@ lean_object* x_543; x_543 = lean_ctor_get(x_542, 0); lean_inc(x_543); lean_dec_ref(x_542); -x_478 = x_538; +x_478 = x_522; x_479 = x_536; -x_480 = x_527; -x_481 = x_529; -x_482 = x_540; -x_483 = x_522; -x_484 = x_526; +x_480 = x_529; +x_481 = x_540; +x_482 = x_526; +x_483 = x_538; +x_484 = x_527; x_485 = x_543; x_486 = lean_box(0); goto block_517; @@ -10354,13 +10350,13 @@ lean_inc(x_547); lean_dec_ref(x_539); x_548 = lean_ctor_get(x_423, 0); lean_inc(x_548); -x_478 = x_538; +x_478 = x_522; x_479 = x_536; -x_480 = x_527; -x_481 = x_529; -x_482 = x_547; -x_483 = x_522; -x_484 = x_526; +x_480 = x_529; +x_481 = x_547; +x_482 = x_526; +x_483 = x_538; +x_484 = x_527; x_485 = x_548; x_486 = lean_box(0); goto block_517; @@ -10737,22 +10733,22 @@ lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; x_620 = lean_ctor_get(x_619, 0); lean_inc(x_620); lean_dec_ref(x_619); -lean_inc_ref(x_607); -lean_inc(x_596); -lean_inc(x_604); +lean_inc_ref(x_600); +lean_inc(x_610); +lean_inc(x_608); x_621 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_621, 0, x_604); -lean_ctor_set(x_621, 1, x_596); -lean_ctor_set(x_621, 2, x_607); +lean_ctor_set(x_621, 0, x_608); +lean_ctor_set(x_621, 1, x_610); +lean_ctor_set(x_621, 2, x_600); lean_inc_ref(x_621); -lean_inc(x_604); -x_622 = l_Lean_Syntax_node1(x_604, x_608, x_621); -lean_inc(x_604); -x_623 = l_Lean_Syntax_node2(x_604, x_598, x_602, x_621); -x_624 = l_Lean_Syntax_node2(x_604, x_609, x_622, x_623); +lean_inc(x_608); +x_622 = l_Lean_Syntax_node1(x_608, x_604, x_621); +lean_inc(x_608); +x_623 = l_Lean_Syntax_node2(x_608, x_596, x_602, x_621); +x_624 = l_Lean_Syntax_node2(x_608, x_597, x_622, x_623); x_625 = lean_unsigned_to_nat(2u); x_626 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_627 = lean_array_push(x_626, x_605); +x_627 = lean_array_push(x_626, x_599); x_628 = lean_array_push(x_627, x_624); x_629 = l_Lake_DSL_expandAttrs(x_7); x_630 = l_Array_append___redArg(x_628, x_629); @@ -10769,21 +10765,21 @@ lean_object* x_634; x_634 = lean_ctor_get(x_633, 0); lean_inc(x_634); lean_dec_ref(x_633); -x_133 = x_614; -x_134 = x_596; -x_135 = x_597; +x_133 = x_598; +x_134 = x_600; +x_135 = x_601; x_136 = x_630; -x_137 = x_625; -x_138 = x_618; -x_139 = x_599; -x_140 = x_600; -x_141 = x_601; -x_142 = x_603; -x_143 = x_620; -x_144 = x_606; -x_145 = x_607; -x_146 = x_610; -x_147 = x_631; +x_137 = x_603; +x_138 = x_620; +x_139 = x_631; +x_140 = x_605; +x_141 = x_607; +x_142 = x_606; +x_143 = x_625; +x_144 = x_614; +x_145 = x_609; +x_146 = x_618; +x_147 = x_610; x_148 = x_634; x_149 = lean_box(0); goto block_230; @@ -10796,15 +10792,15 @@ lean_dec_ref(x_630); lean_dec(x_620); lean_dec(x_618); lean_dec(x_614); -lean_dec_ref(x_610); -lean_dec_ref(x_607); -lean_dec(x_606); +lean_dec(x_610); +lean_dec_ref(x_609); +lean_dec(x_607); +lean_dec_ref(x_606); +lean_dec_ref(x_605); lean_dec(x_603); -lean_dec_ref(x_601); -lean_dec(x_600); -lean_dec(x_599); -lean_dec_ref(x_597); -lean_dec(x_596); +lean_dec(x_601); +lean_dec_ref(x_600); +lean_dec(x_598); lean_dec(x_6); x_635 = lean_ctor_get(x_633, 0); lean_inc(x_635); @@ -10834,21 +10830,21 @@ lean_dec_ref(x_10); x_638 = lean_ctor_get(x_594, 0); lean_inc(x_638); lean_dec_ref(x_594); -x_133 = x_614; -x_134 = x_596; -x_135 = x_597; +x_133 = x_598; +x_134 = x_600; +x_135 = x_601; x_136 = x_630; -x_137 = x_625; -x_138 = x_618; -x_139 = x_599; -x_140 = x_600; -x_141 = x_601; -x_142 = x_603; -x_143 = x_620; -x_144 = x_606; -x_145 = x_607; -x_146 = x_610; -x_147 = x_631; +x_137 = x_603; +x_138 = x_620; +x_139 = x_631; +x_140 = x_605; +x_141 = x_607; +x_142 = x_606; +x_143 = x_625; +x_144 = x_614; +x_145 = x_609; +x_146 = x_618; +x_147 = x_610; x_148 = x_638; x_149 = lean_box(0); goto block_230; @@ -10860,20 +10856,20 @@ lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_dec(x_618); lean_dec(x_616); lean_dec(x_614); -lean_dec_ref(x_610); -lean_dec(x_609); +lean_dec(x_610); +lean_dec_ref(x_609); lean_dec(x_608); -lean_dec_ref(x_607); -lean_dec(x_606); -lean_dec(x_605); +lean_dec(x_607); +lean_dec_ref(x_606); +lean_dec_ref(x_605); lean_dec(x_604); lean_dec(x_603); lean_dec(x_602); -lean_dec_ref(x_601); -lean_dec(x_600); +lean_dec(x_601); +lean_dec_ref(x_600); lean_dec(x_599); lean_dec(x_598); -lean_dec_ref(x_597); +lean_dec(x_597); lean_dec(x_596); lean_dec_ref(x_587); lean_dec(x_11); @@ -10904,20 +10900,20 @@ else lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_dec(x_616); lean_dec(x_614); -lean_dec_ref(x_610); -lean_dec(x_609); +lean_dec(x_610); +lean_dec_ref(x_609); lean_dec(x_608); -lean_dec_ref(x_607); -lean_dec(x_606); -lean_dec(x_605); +lean_dec(x_607); +lean_dec_ref(x_606); +lean_dec_ref(x_605); lean_dec(x_604); lean_dec(x_603); lean_dec(x_602); -lean_dec_ref(x_601); -lean_dec(x_600); +lean_dec(x_601); +lean_dec_ref(x_600); lean_dec(x_599); lean_dec(x_598); -lean_dec_ref(x_597); +lean_dec(x_597); lean_dec(x_596); lean_dec_ref(x_587); lean_dec(x_11); @@ -10947,20 +10943,20 @@ else { lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_dec(x_614); -lean_dec_ref(x_610); -lean_dec(x_609); +lean_dec(x_610); +lean_dec_ref(x_609); lean_dec(x_608); -lean_dec_ref(x_607); -lean_dec(x_606); -lean_dec(x_605); +lean_dec(x_607); +lean_dec_ref(x_606); +lean_dec_ref(x_605); lean_dec(x_604); lean_dec(x_603); lean_dec(x_602); -lean_dec_ref(x_601); -lean_dec(x_600); +lean_dec(x_601); +lean_dec_ref(x_600); lean_dec(x_599); lean_dec(x_598); -lean_dec_ref(x_597); +lean_dec(x_597); lean_dec(x_596); lean_dec_ref(x_587); lean_dec(x_11); @@ -10988,20 +10984,20 @@ return x_647; } else { -lean_dec_ref(x_610); -lean_dec(x_609); +lean_dec(x_610); +lean_dec_ref(x_609); lean_dec(x_608); -lean_dec_ref(x_607); -lean_dec(x_606); -lean_dec(x_605); +lean_dec(x_607); +lean_dec_ref(x_606); +lean_dec_ref(x_605); lean_dec(x_604); lean_dec(x_603); lean_dec(x_602); -lean_dec_ref(x_601); -lean_dec(x_600); +lean_dec(x_601); +lean_dec_ref(x_600); lean_dec(x_599); lean_dec(x_598); -lean_dec_ref(x_597); +lean_dec(x_597); lean_dec(x_596); lean_dec_ref(x_587); lean_dec(x_11); @@ -11034,28 +11030,28 @@ x_664 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_665 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_666 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_667 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_649); +lean_inc(x_654); x_668 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_668, 0, x_649); +lean_ctor_set(x_668, 0, x_654); lean_ctor_set(x_668, 1, x_666); lean_ctor_set(x_668, 2, x_667); lean_inc_ref(x_668); -lean_inc(x_649); -x_669 = l_Lean_Syntax_node1(x_649, x_665, x_668); +lean_inc(x_654); +x_669 = l_Lean_Syntax_node1(x_654, x_665, x_668); x_670 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_671 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_672 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_673 = l_Lean_addMacroScope(x_656, x_672, x_653); -lean_inc(x_655); -lean_inc(x_649); +x_673 = l_Lean_addMacroScope(x_656, x_672, x_652); +lean_inc(x_653); +lean_inc(x_654); x_674 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_674, 0, x_649); +lean_ctor_set(x_674, 0, x_654); lean_ctor_set(x_674, 1, x_671); lean_ctor_set(x_674, 2, x_673); -lean_ctor_set(x_674, 3, x_655); -lean_inc(x_649); -x_675 = l_Lean_Syntax_node2(x_649, x_670, x_674, x_668); -x_676 = l_Lean_Syntax_node2(x_649, x_664, x_669, x_675); +lean_ctor_set(x_674, 3, x_653); +lean_inc(x_654); +x_675 = l_Lean_Syntax_node2(x_654, x_670, x_674, x_668); +x_676 = l_Lean_Syntax_node2(x_654, x_664, x_669, x_675); if (lean_obj_tag(x_594) == 0) { lean_object* x_677; lean_object* x_678; @@ -11067,21 +11063,21 @@ x_678 = lean_apply_3(x_677, x_10, x_11, lean_box(0)); if (lean_obj_tag(x_678) == 0) { lean_dec_ref(x_678); -x_596 = x_666; -x_597 = x_661; -x_598 = x_670; -x_599 = x_654; -x_600 = x_655; -x_601 = x_662; +x_596 = x_670; +x_597 = x_664; +x_598 = x_651; +x_599 = x_676; +x_600 = x_667; +x_601 = x_655; x_602 = x_650; -x_603 = x_651; -x_604 = x_659; -x_605 = x_676; -x_606 = x_652; -x_607 = x_667; -x_608 = x_665; -x_609 = x_664; -x_610 = x_663; +x_603 = x_649; +x_604 = x_665; +x_605 = x_662; +x_606 = x_663; +x_607 = x_653; +x_608 = x_659; +x_609 = x_661; +x_610 = x_666; x_611 = lean_box(0); goto block_648; } @@ -11091,10 +11087,10 @@ lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_dec(x_676); lean_dec(x_659); lean_dec(x_655); -lean_dec(x_654); -lean_dec(x_652); +lean_dec(x_653); lean_dec(x_651); lean_dec(x_650); +lean_dec(x_649); lean_dec_ref(x_587); lean_dec(x_11); lean_dec_ref(x_10); @@ -11122,21 +11118,21 @@ return x_681; } else { -x_596 = x_666; -x_597 = x_661; -x_598 = x_670; -x_599 = x_654; -x_600 = x_655; -x_601 = x_662; +x_596 = x_670; +x_597 = x_664; +x_598 = x_651; +x_599 = x_676; +x_600 = x_667; +x_601 = x_655; x_602 = x_650; -x_603 = x_651; -x_604 = x_659; -x_605 = x_676; -x_606 = x_652; -x_607 = x_667; -x_608 = x_665; -x_609 = x_664; -x_610 = x_663; +x_603 = x_649; +x_604 = x_665; +x_605 = x_662; +x_606 = x_663; +x_607 = x_653; +x_608 = x_659; +x_609 = x_661; +x_610 = x_666; x_611 = lean_box(0); goto block_648; } @@ -11289,13 +11285,13 @@ lean_object* x_714; x_714 = lean_ctor_get(x_713, 0); lean_inc(x_714); lean_dec_ref(x_713); -x_649 = x_709; +x_649 = x_693; x_650 = x_707; -x_651 = x_698; -x_652 = x_700; -x_653 = x_711; -x_654 = x_693; -x_655 = x_697; +x_651 = x_700; +x_652 = x_711; +x_653 = x_697; +x_654 = x_709; +x_655 = x_698; x_656 = x_714; x_657 = lean_box(0); goto block_688; @@ -11342,13 +11338,13 @@ lean_inc(x_718); lean_dec_ref(x_710); x_719 = lean_ctor_get(x_594, 0); lean_inc(x_719); -x_649 = x_709; +x_649 = x_693; x_650 = x_707; -x_651 = x_698; -x_652 = x_700; -x_653 = x_718; -x_654 = x_693; -x_655 = x_697; +x_651 = x_700; +x_652 = x_718; +x_653 = x_697; +x_654 = x_709; +x_655 = x_698; x_656 = x_719; x_657 = lean_box(0); goto block_688; @@ -11556,43 +11552,43 @@ return x_742; block_132: { lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -lean_inc_ref(x_22); -x_44 = l_Array_append___redArg(x_22, x_43); +lean_inc_ref(x_16); +x_44 = l_Array_append___redArg(x_16, x_43); lean_dec_ref(x_43); +lean_inc(x_42); lean_inc(x_26); -lean_inc(x_19); x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_19); -lean_ctor_set(x_45, 1, x_26); +lean_ctor_set(x_45, 0, x_26); +lean_ctor_set(x_45, 1, x_42); lean_ctor_set(x_45, 2, x_44); -lean_inc_n(x_25, 6); -lean_inc(x_30); -lean_inc(x_19); -x_46 = l_Lean_Syntax_node7(x_19, x_30, x_45, x_25, x_25, x_25, x_25, x_25, x_25); +lean_inc_n(x_22, 6); +lean_inc(x_41); +lean_inc(x_26); +x_46 = l_Lean_Syntax_node7(x_26, x_41, x_45, x_22, x_22, x_22, x_22, x_22, x_22); x_47 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; -lean_inc_ref(x_34); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_48 = l_Lean_Name_mkStr4(x_15, x_33, x_34, x_47); -lean_inc(x_19); +lean_inc_ref(x_30); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_48 = l_Lean_Name_mkStr4(x_40, x_35, x_30, x_47); +lean_inc(x_26); x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_19); +lean_ctor_set(x_49, 0, x_26); lean_ctor_set(x_49, 1, x_47); x_50 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; -lean_inc_ref(x_34); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_51 = l_Lean_Name_mkStr4(x_15, x_33, x_34, x_50); +lean_inc_ref(x_30); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_51 = l_Lean_Name_mkStr4(x_40, x_35, x_30, x_50); x_52 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_53 = lean_box(2); -lean_inc(x_26); +lean_inc(x_42); x_54 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_26); +lean_ctor_set(x_54, 1, x_42); lean_ctor_set(x_54, 2, x_52); -x_55 = lean_mk_empty_array_with_capacity(x_17); -lean_inc(x_31); -x_56 = lean_array_push(x_55, x_31); +x_55 = lean_mk_empty_array_with_capacity(x_39); +lean_inc(x_20); +x_56 = lean_array_push(x_55, x_20); x_57 = lean_array_push(x_56, x_54); lean_inc(x_51); x_58 = lean_alloc_ctor(1, 3, 0); @@ -11600,194 +11596,194 @@ lean_ctor_set(x_58, 0, x_53); lean_ctor_set(x_58, 1, x_51); lean_ctor_set(x_58, 2, x_57); x_59 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_34); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_60 = l_Lean_Name_mkStr4(x_15, x_33, x_34, x_59); +lean_inc_ref(x_30); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_60 = l_Lean_Name_mkStr4(x_40, x_35, x_30, x_59); x_61 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_42); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_62 = l_Lean_Name_mkStr4(x_15, x_33, x_42, x_61); -lean_inc(x_37); +lean_inc_ref(x_36); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_62 = l_Lean_Name_mkStr4(x_40, x_35, x_36, x_61); +lean_inc(x_33); lean_inc(x_62); -lean_inc(x_19); -x_63 = l_Lean_Syntax_node2(x_19, x_62, x_37, x_14); lean_inc(x_26); -lean_inc(x_19); -x_64 = l_Lean_Syntax_node1(x_19, x_26, x_63); -lean_inc(x_25); +x_63 = l_Lean_Syntax_node2(x_26, x_62, x_33, x_38); +lean_inc(x_42); +lean_inc(x_26); +x_64 = l_Lean_Syntax_node1(x_26, x_42, x_63); +lean_inc(x_22); lean_inc(x_60); -lean_inc(x_19); -x_65 = l_Lean_Syntax_node2(x_19, x_60, x_25, x_64); +lean_inc(x_26); +x_65 = l_Lean_Syntax_node2(x_26, x_60, x_22, x_64); x_66 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__60; -lean_inc_ref(x_34); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_67 = l_Lean_Name_mkStr4(x_15, x_33, x_34, x_66); +lean_inc_ref(x_30); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_67 = l_Lean_Name_mkStr4(x_40, x_35, x_30, x_66); x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__1; x_69 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__2; -lean_inc_ref(x_38); -x_70 = l_Lean_Name_mkStr3(x_38, x_13, x_69); -lean_inc(x_35); +lean_inc_ref(x_15); +x_70 = l_Lean_Name_mkStr3(x_15, x_24, x_69); +lean_inc(x_21); lean_inc(x_70); -lean_inc(x_28); -x_71 = l_Lean_addMacroScope(x_28, x_70, x_35); -lean_inc(x_39); +lean_inc(x_23); +x_71 = l_Lean_addMacroScope(x_23, x_70, x_21); +lean_inc(x_31); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_39); -lean_inc(x_20); +lean_ctor_set(x_72, 1, x_31); +lean_inc(x_37); x_73 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_20); -lean_inc(x_19); +lean_ctor_set(x_73, 1, x_37); +lean_inc(x_26); x_74 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_74, 0, x_19); +lean_ctor_set(x_74, 0, x_26); lean_ctor_set(x_74, 1, x_68); lean_ctor_set(x_74, 2, x_71); lean_ctor_set(x_74, 3, x_73); +lean_inc(x_42); lean_inc(x_26); -lean_inc(x_19); -x_75 = l_Lean_Syntax_node4(x_19, x_26, x_23, x_40, x_24, x_36); -lean_inc(x_19); -x_76 = l_Lean_Syntax_node2(x_19, x_18, x_74, x_75); +x_75 = l_Lean_Syntax_node4(x_26, x_42, x_28, x_14, x_34, x_18); +lean_inc(x_26); +x_76 = l_Lean_Syntax_node2(x_26, x_25, x_74, x_75); x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__62; x_78 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__63; -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_79 = l_Lean_Name_mkStr4(x_15, x_33, x_77, x_78); -lean_inc_n(x_25, 2); -lean_inc(x_19); -x_80 = l_Lean_Syntax_node2(x_19, x_79, x_25, x_25); -lean_inc(x_25); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_79 = l_Lean_Name_mkStr4(x_40, x_35, x_77, x_78); +lean_inc_n(x_22, 2); +lean_inc(x_26); +x_80 = l_Lean_Syntax_node2(x_26, x_79, x_22, x_22); +lean_inc(x_22); lean_inc(x_80); -lean_inc(x_32); +lean_inc(x_29); lean_inc(x_67); -lean_inc(x_19); -x_81 = l_Lean_Syntax_node4(x_19, x_67, x_32, x_76, x_80, x_25); -lean_inc(x_19); -x_82 = l_Lean_Syntax_node4(x_19, x_48, x_49, x_58, x_65, x_81); +lean_inc(x_26); +x_81 = l_Lean_Syntax_node4(x_26, x_67, x_29, x_76, x_80, x_22); +lean_inc(x_26); +x_82 = l_Lean_Syntax_node4(x_26, x_48, x_49, x_58, x_65, x_81); lean_inc(x_27); -lean_inc(x_19); -x_83 = l_Lean_Syntax_node2(x_19, x_27, x_46, x_82); +lean_inc(x_26); +x_83 = l_Lean_Syntax_node2(x_26, x_27, x_46, x_82); x_84 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_42); -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_85 = l_Lean_Name_mkStr4(x_15, x_33, x_42, x_84); +lean_inc_ref(x_36); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_85 = l_Lean_Name_mkStr4(x_40, x_35, x_36, x_84); x_86 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_19); +lean_inc(x_26); x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_19); +lean_ctor_set(x_87, 0, x_26); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Syntax_SepArray_ofElems(x_41, x_16); -lean_dec_ref(x_16); -x_89 = l_Array_append___redArg(x_22, x_88); +x_88 = l_Lean_Syntax_SepArray_ofElems(x_17, x_32); +lean_dec_ref(x_32); +x_89 = l_Array_append___redArg(x_16, x_88); lean_dec_ref(x_88); +lean_inc(x_42); lean_inc(x_26); -lean_inc(x_19); x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_19); -lean_ctor_set(x_90, 1, x_26); +lean_ctor_set(x_90, 0, x_26); +lean_ctor_set(x_90, 1, x_42); lean_ctor_set(x_90, 2, x_89); x_91 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_19); +lean_inc(x_26); x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_19); +lean_ctor_set(x_92, 0, x_26); lean_ctor_set(x_92, 1, x_91); -lean_inc(x_19); -x_93 = l_Lean_Syntax_node3(x_19, x_85, x_87, x_90, x_92); lean_inc(x_26); -lean_inc(x_19); -x_94 = l_Lean_Syntax_node1(x_19, x_26, x_93); -lean_inc_n(x_25, 6); -lean_inc(x_19); -x_95 = l_Lean_Syntax_node7(x_19, x_30, x_25, x_94, x_25, x_25, x_25, x_25, x_25); +x_93 = l_Lean_Syntax_node3(x_26, x_85, x_87, x_90, x_92); +lean_inc(x_42); +lean_inc(x_26); +x_94 = l_Lean_Syntax_node1(x_26, x_42, x_93); +lean_inc_n(x_22, 6); +lean_inc(x_26); +x_95 = l_Lean_Syntax_node7(x_26, x_41, x_22, x_94, x_22, x_22, x_22, x_22, x_22); x_96 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; -lean_inc_ref(x_33); -lean_inc_ref(x_15); -x_97 = l_Lean_Name_mkStr4(x_15, x_33, x_34, x_96); +lean_inc_ref(x_35); +lean_inc_ref(x_40); +x_97 = l_Lean_Name_mkStr4(x_40, x_35, x_30, x_96); x_98 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_19); +lean_inc(x_26); x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_19); +lean_ctor_set(x_99, 0, x_26); lean_ctor_set(x_99, 1, x_98); x_100 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_101 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_35); -lean_inc(x_28); -x_102 = l_Lean_addMacroScope(x_28, x_101, x_35); -lean_inc(x_20); -lean_inc(x_19); +lean_inc(x_21); +lean_inc(x_23); +x_102 = l_Lean_addMacroScope(x_23, x_101, x_21); +lean_inc(x_37); +lean_inc(x_26); x_103 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_103, 0, x_19); +lean_ctor_set(x_103, 0, x_26); lean_ctor_set(x_103, 1, x_100); lean_ctor_set(x_103, 2, x_102); -lean_ctor_set(x_103, 3, x_20); -lean_inc(x_25); -lean_inc(x_19); -x_104 = l_Lean_Syntax_node2(x_19, x_51, x_103, x_25); +lean_ctor_set(x_103, 3, x_37); +lean_inc(x_22); +lean_inc(x_26); +x_104 = l_Lean_Syntax_node2(x_26, x_51, x_103, x_22); x_105 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__5; x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_107 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_35); -lean_inc(x_28); -x_108 = l_Lean_addMacroScope(x_28, x_107, x_35); -x_109 = l_Lean_Name_mkStr2(x_38, x_105); +lean_inc(x_21); +lean_inc(x_23); +x_108 = l_Lean_addMacroScope(x_23, x_107, x_21); +x_109 = l_Lean_Name_mkStr2(x_15, x_105); lean_inc(x_109); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_39); +lean_ctor_set(x_110, 1, x_31); x_111 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_111, 0, x_109); -lean_inc(x_20); +lean_inc(x_37); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_20); +lean_ctor_set(x_112, 1, x_37); x_113 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_113, 0, x_110); lean_ctor_set(x_113, 1, x_112); -lean_inc(x_19); +lean_inc(x_26); x_114 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_114, 0, x_19); +lean_ctor_set(x_114, 0, x_26); lean_ctor_set(x_114, 1, x_106); lean_ctor_set(x_114, 2, x_108); lean_ctor_set(x_114, 3, x_113); -lean_inc(x_19); -x_115 = l_Lean_Syntax_node2(x_19, x_62, x_37, x_114); lean_inc(x_26); -lean_inc(x_19); -x_116 = l_Lean_Syntax_node1(x_19, x_26, x_115); -lean_inc(x_25); -lean_inc(x_19); -x_117 = l_Lean_Syntax_node2(x_19, x_60, x_25, x_116); +x_115 = l_Lean_Syntax_node2(x_26, x_62, x_33, x_114); +lean_inc(x_42); +lean_inc(x_26); +x_116 = l_Lean_Syntax_node1(x_26, x_42, x_115); +lean_inc(x_22); +lean_inc(x_26); +x_117 = l_Lean_Syntax_node2(x_26, x_60, x_22, x_116); x_118 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_119 = l_Lean_Name_mkStr4(x_15, x_33, x_42, x_118); +x_119 = l_Lean_Name_mkStr4(x_40, x_35, x_36, x_118); x_120 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_19); +lean_inc(x_26); x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_19); +lean_ctor_set(x_121, 0, x_26); lean_ctor_set(x_121, 1, x_120); x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_123 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_124 = l_Lean_addMacroScope(x_28, x_123, x_35); -lean_inc(x_19); +x_124 = l_Lean_addMacroScope(x_23, x_123, x_21); +lean_inc(x_26); x_125 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_125, 0, x_19); +lean_ctor_set(x_125, 0, x_26); lean_ctor_set(x_125, 1, x_122); lean_ctor_set(x_125, 2, x_124); -lean_ctor_set(x_125, 3, x_20); -lean_inc(x_19); -x_126 = l_Lean_Syntax_node3(x_19, x_119, x_31, x_121, x_125); -lean_inc(x_25); -lean_inc(x_19); -x_127 = l_Lean_Syntax_node4(x_19, x_67, x_32, x_126, x_80, x_25); -lean_inc(x_19); -x_128 = l_Lean_Syntax_node5(x_19, x_97, x_99, x_104, x_117, x_127, x_25); -lean_inc(x_19); -x_129 = l_Lean_Syntax_node2(x_19, x_27, x_95, x_128); -x_130 = l_Lean_Syntax_node3(x_19, x_26, x_21, x_83, x_129); +lean_ctor_set(x_125, 3, x_37); +lean_inc(x_26); +x_126 = l_Lean_Syntax_node3(x_26, x_119, x_20, x_121, x_125); +lean_inc(x_22); +lean_inc(x_26); +x_127 = l_Lean_Syntax_node4(x_26, x_67, x_29, x_126, x_80, x_22); +lean_inc(x_26); +x_128 = l_Lean_Syntax_node5(x_26, x_97, x_99, x_104, x_117, x_127, x_22); +lean_inc(x_26); +x_129 = l_Lean_Syntax_node2(x_26, x_27, x_95, x_128); +x_130 = l_Lean_Syntax_node3(x_26, x_42, x_13, x_83, x_129); x_131 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_131, 0, x_130); return x_131; @@ -11797,91 +11793,91 @@ return x_131; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__0; x_151 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; -lean_inc_ref(x_145); -lean_inc(x_134); -lean_inc(x_138); +lean_inc_ref(x_134); +lean_inc(x_147); +lean_inc(x_146); x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_138); -lean_ctor_set(x_152, 1, x_134); -lean_ctor_set(x_152, 2, x_145); +lean_ctor_set(x_152, 0, x_146); +lean_ctor_set(x_152, 1, x_147); +lean_ctor_set(x_152, 2, x_134); x_153 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; -lean_inc(x_138); +lean_inc(x_146); x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_138); +lean_ctor_set(x_154, 0, x_146); lean_ctor_set(x_154, 1, x_153); x_155 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_138); +lean_inc(x_146); x_156 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_156, 0, x_138); +lean_ctor_set(x_156, 0, x_146); lean_ctor_set(x_156, 1, x_155); x_157 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_158 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_143); +lean_inc(x_138); lean_inc(x_148); -x_159 = l_Lean_addMacroScope(x_148, x_158, x_143); +x_159 = l_Lean_addMacroScope(x_148, x_158, x_138); x_160 = lean_box(0); x_161 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__3; -lean_inc(x_140); +lean_inc(x_141); x_162 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_140); -lean_inc(x_138); +lean_ctor_set(x_162, 1, x_141); +lean_inc(x_146); x_163 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_163, 0, x_138); +lean_ctor_set(x_163, 0, x_146); lean_ctor_set(x_163, 1, x_157); lean_ctor_set(x_163, 2, x_159); lean_ctor_set(x_163, 3, x_162); x_164 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_146); -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_165 = l_Lean_Name_mkStr4(x_135, x_141, x_146, x_164); +lean_inc_ref(x_142); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_165 = l_Lean_Name_mkStr4(x_145, x_140, x_142, x_164); x_166 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_146); -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_167 = l_Lean_Name_mkStr4(x_135, x_141, x_146, x_166); +lean_inc_ref(x_142); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_167 = l_Lean_Name_mkStr4(x_145, x_140, x_142, x_166); x_168 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_138); +lean_inc(x_146); x_169 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_169, 0, x_138); +lean_ctor_set(x_169, 0, x_146); lean_ctor_set(x_169, 1, x_168); x_170 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_171 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_172 = lean_box(0); -lean_inc(x_143); +lean_inc(x_138); lean_inc(x_148); -x_173 = l_Lean_addMacroScope(x_148, x_172, x_143); +x_173 = l_Lean_addMacroScope(x_148, x_172, x_138); x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__1; x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; x_176 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__59; -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_177 = l_Lean_Name_mkStr3(x_135, x_141, x_176); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_177 = l_Lean_Name_mkStr3(x_145, x_140, x_176); x_178 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_178, 0, x_177); x_179 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_135); -x_180 = l_Lean_Name_mkStr3(x_135, x_179, x_176); +lean_inc_ref(x_145); +x_180 = l_Lean_Name_mkStr3(x_145, x_179, x_176); x_181 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_181, 0, x_180); -lean_inc_ref(x_135); -x_182 = l_Lean_Name_mkStr2(x_135, x_179); +lean_inc_ref(x_145); +x_182 = l_Lean_Name_mkStr2(x_145, x_179); x_183 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_183, 0, x_182); -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_184 = l_Lean_Name_mkStr2(x_135, x_141); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_184 = l_Lean_Name_mkStr2(x_145, x_140); x_185 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_185, 0, x_184); -lean_inc_ref(x_135); -x_186 = l_Lean_Name_mkStr1(x_135); +lean_inc_ref(x_145); +x_186 = l_Lean_Name_mkStr1(x_145); x_187 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_187, 0, x_186); -lean_inc(x_140); +lean_inc(x_141); x_188 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_140); +lean_ctor_set(x_188, 1, x_141); x_189 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_189, 0, x_185); lean_ctor_set(x_189, 1, x_188); @@ -11897,129 +11893,129 @@ lean_ctor_set(x_192, 1, x_191); x_193 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_193, 0, x_175); lean_ctor_set(x_193, 1, x_192); -lean_inc(x_138); +lean_inc(x_146); x_194 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_194, 0, x_138); +lean_ctor_set(x_194, 0, x_146); lean_ctor_set(x_194, 1, x_171); lean_ctor_set(x_194, 2, x_173); lean_ctor_set(x_194, 3, x_193); -lean_inc(x_138); -x_195 = l_Lean_Syntax_node1(x_138, x_170, x_194); -lean_inc(x_138); -x_196 = l_Lean_Syntax_node2(x_138, x_167, x_169, x_195); +lean_inc(x_146); +x_195 = l_Lean_Syntax_node1(x_146, x_170, x_194); +lean_inc(x_146); +x_196 = l_Lean_Syntax_node2(x_146, x_167, x_169, x_195); x_197 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__37; x_198 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__38; -lean_inc(x_138); +lean_inc(x_146); x_199 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_199, 0, x_138); +lean_ctor_set(x_199, 0, x_146); lean_ctor_set(x_199, 1, x_198); -lean_inc(x_138); -x_200 = l_Lean_Syntax_node1(x_138, x_197, x_199); +lean_inc(x_146); +x_200 = l_Lean_Syntax_node1(x_146, x_197, x_199); x_201 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -lean_inc(x_138); +lean_inc(x_146); x_202 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_202, 0, x_138); +lean_ctor_set(x_202, 0, x_146); lean_ctor_set(x_202, 1, x_201); -lean_inc(x_144); -lean_inc(x_134); -lean_inc(x_138); -x_203 = l_Lean_Syntax_node1(x_138, x_134, x_144); +lean_inc(x_133); +lean_inc(x_147); +lean_inc(x_146); +x_203 = l_Lean_Syntax_node1(x_146, x_147, x_133); lean_inc(x_200); -lean_inc(x_134); -lean_inc(x_138); -x_204 = l_Lean_Syntax_node3(x_138, x_134, x_200, x_202, x_203); +lean_inc(x_147); +lean_inc(x_146); +x_204 = l_Lean_Syntax_node3(x_146, x_147, x_200, x_202, x_203); x_205 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_138); +lean_inc(x_146); x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_138); +lean_ctor_set(x_206, 0, x_146); lean_ctor_set(x_206, 1, x_205); -lean_inc(x_138); -x_207 = l_Lean_Syntax_node3(x_138, x_165, x_196, x_204, x_206); +lean_inc(x_146); +x_207 = l_Lean_Syntax_node3(x_146, x_165, x_196, x_204, x_206); x_208 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_138); +lean_inc(x_146); x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_138); +lean_ctor_set(x_209, 0, x_146); lean_ctor_set(x_209, 1, x_208); x_210 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_146); -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_211 = l_Lean_Name_mkStr4(x_135, x_141, x_146, x_210); +lean_inc_ref(x_142); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_211 = l_Lean_Name_mkStr4(x_145, x_140, x_142, x_210); x_212 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_213 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_143); +lean_inc(x_138); lean_inc(x_148); -x_214 = l_Lean_addMacroScope(x_148, x_213, x_143); +x_214 = l_Lean_addMacroScope(x_148, x_213, x_138); x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__8; x_216 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_140); +lean_inc(x_141); x_217 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_217, 0, x_216); -lean_ctor_set(x_217, 1, x_140); +lean_ctor_set(x_217, 1, x_141); x_218 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_218, 0, x_215); lean_ctor_set(x_218, 1, x_217); -lean_inc(x_138); +lean_inc(x_146); x_219 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_219, 0, x_138); +lean_ctor_set(x_219, 0, x_146); lean_ctor_set(x_219, 1, x_212); lean_ctor_set(x_219, 2, x_214); lean_ctor_set(x_219, 3, x_218); +lean_inc(x_139); lean_inc(x_147); -lean_inc(x_134); -lean_inc(x_138); -x_220 = l_Lean_Syntax_node1(x_138, x_134, x_147); +lean_inc(x_146); +x_220 = l_Lean_Syntax_node1(x_146, x_147, x_139); lean_inc(x_211); -lean_inc(x_138); -x_221 = l_Lean_Syntax_node2(x_138, x_211, x_219, x_220); +lean_inc(x_146); +x_221 = l_Lean_Syntax_node2(x_146, x_211, x_219, x_220); lean_inc_ref(x_209); lean_inc_ref(x_156); -lean_inc(x_139); +lean_inc(x_137); lean_inc_ref(x_152); -lean_inc(x_138); -x_222 = l_Lean_Syntax_node8(x_138, x_151, x_152, x_154, x_139, x_156, x_163, x_207, x_209, x_221); +lean_inc(x_146); +x_222 = l_Lean_Syntax_node8(x_146, x_151, x_152, x_154, x_137, x_156, x_163, x_207, x_209, x_221); x_223 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_224 = l_Lean_Name_mkStr4(x_135, x_141, x_176, x_223); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_224 = l_Lean_Name_mkStr4(x_145, x_140, x_176, x_223); x_225 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_141); -lean_inc_ref(x_135); -x_226 = l_Lean_Name_mkStr4(x_135, x_141, x_176, x_225); +lean_inc_ref(x_140); +lean_inc_ref(x_145); +x_226 = l_Lean_Name_mkStr4(x_145, x_140, x_176, x_225); if (lean_obj_tag(x_6) == 0) { lean_object* x_227; x_227 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_13 = x_174; +x_13 = x_222; x_14 = x_133; -x_15 = x_135; -x_16 = x_136; -x_17 = x_137; -x_18 = x_211; -x_19 = x_138; -x_20 = x_140; -x_21 = x_222; -x_22 = x_145; -x_23 = x_200; -x_24 = x_147; -x_25 = x_152; -x_26 = x_134; +x_15 = x_150; +x_16 = x_134; +x_17 = x_201; +x_18 = x_135; +x_19 = lean_box(0); +x_20 = x_137; +x_21 = x_138; +x_22 = x_152; +x_23 = x_148; +x_24 = x_174; +x_25 = x_211; +x_26 = x_146; x_27 = x_224; -x_28 = x_148; -x_29 = lean_box(0); -x_30 = x_226; -x_31 = x_139; -x_32 = x_209; -x_33 = x_141; -x_34 = x_176; -x_35 = x_143; +x_28 = x_200; +x_29 = x_209; +x_30 = x_176; +x_31 = x_160; +x_32 = x_136; +x_33 = x_156; +x_34 = x_139; +x_35 = x_140; x_36 = x_142; -x_37 = x_156; -x_38 = x_150; -x_39 = x_160; -x_40 = x_144; -x_41 = x_201; -x_42 = x_146; +x_37 = x_141; +x_38 = x_144; +x_39 = x_143; +x_40 = x_145; +x_41 = x_226; +x_42 = x_147; x_43 = x_227; goto block_132; } @@ -12030,36 +12026,36 @@ x_228 = lean_ctor_get(x_6, 0); lean_inc(x_228); lean_dec_ref(x_6); x_229 = l_Array_mkArray1___redArg(x_228); -x_13 = x_174; +x_13 = x_222; x_14 = x_133; -x_15 = x_135; -x_16 = x_136; -x_17 = x_137; -x_18 = x_211; -x_19 = x_138; -x_20 = x_140; -x_21 = x_222; -x_22 = x_145; -x_23 = x_200; -x_24 = x_147; -x_25 = x_152; -x_26 = x_134; +x_15 = x_150; +x_16 = x_134; +x_17 = x_201; +x_18 = x_135; +x_19 = lean_box(0); +x_20 = x_137; +x_21 = x_138; +x_22 = x_152; +x_23 = x_148; +x_24 = x_174; +x_25 = x_211; +x_26 = x_146; x_27 = x_224; -x_28 = x_148; -x_29 = lean_box(0); -x_30 = x_226; -x_31 = x_139; -x_32 = x_209; -x_33 = x_141; -x_34 = x_176; -x_35 = x_143; +x_28 = x_200; +x_29 = x_209; +x_30 = x_176; +x_31 = x_160; +x_32 = x_136; +x_33 = x_156; +x_34 = x_139; +x_35 = x_140; x_36 = x_142; -x_37 = x_156; -x_38 = x_150; -x_39 = x_160; -x_40 = x_144; -x_41 = x_201; -x_42 = x_146; +x_37 = x_141; +x_38 = x_144; +x_39 = x_143; +x_40 = x_145; +x_41 = x_226; +x_42 = x_147; x_43 = x_229; goto block_132; } @@ -12898,56 +12894,56 @@ x_22 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkC x_23 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__2; x_24 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_25 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_20); +lean_inc(x_19); if (lean_is_scalar(x_13)) { x_26 = lean_alloc_ctor(1, 3, 0); } else { x_26 = x_13; lean_ctor_set_tag(x_26, 1); } -lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 0, x_19); lean_ctor_set(x_26, 1, x_24); lean_ctor_set(x_26, 2, x_25); lean_inc_ref_n(x_26, 7); -lean_inc(x_20); -x_27 = l_Lean_Syntax_node7(x_20, x_23, x_26, x_26, x_26, x_26, x_26, x_26, x_26); +lean_inc(x_19); +x_27 = l_Lean_Syntax_node7(x_19, x_23, x_26, x_26, x_26, x_26, x_26, x_26, x_26); x_28 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__3; x_29 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_20); +lean_inc(x_19); x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_20); +lean_ctor_set(x_30, 0, x_19); lean_ctor_set(x_30, 1, x_29); x_31 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__4; x_32 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; -lean_inc(x_18); +lean_inc(x_20); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_18); +lean_ctor_set(x_33, 0, x_20); lean_ctor_set(x_33, 1, x_24); lean_ctor_set(x_33, 2, x_32); x_34 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; x_35 = lean_array_push(x_34, x_3); x_36 = lean_array_push(x_35, x_33); x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_18); +lean_ctor_set(x_37, 0, x_20); lean_ctor_set(x_37, 1, x_31); lean_ctor_set(x_37, 2, x_36); x_38 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__5; x_39 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__58; x_40 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_20); +lean_inc(x_19); x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_20); +lean_ctor_set(x_41, 0, x_19); lean_ctor_set(x_41, 1, x_40); -lean_inc(x_20); -x_42 = l_Lean_Syntax_node2(x_20, x_39, x_41, x_4); -lean_inc(x_20); -x_43 = l_Lean_Syntax_node1(x_20, x_24, x_42); +lean_inc(x_19); +x_42 = l_Lean_Syntax_node2(x_19, x_39, x_41, x_4); +lean_inc(x_19); +x_43 = l_Lean_Syntax_node1(x_19, x_24, x_42); lean_inc_ref(x_26); -lean_inc(x_20); -x_44 = l_Lean_Syntax_node2(x_20, x_38, x_26, x_43); -lean_inc(x_20); -x_45 = l_Lean_Syntax_node5(x_20, x_28, x_30, x_37, x_44, x_19, x_26); -x_46 = l_Lean_Syntax_node2(x_20, x_22, x_27, x_45); +lean_inc(x_19); +x_44 = l_Lean_Syntax_node2(x_19, x_38, x_26, x_43); +lean_inc(x_19); +x_45 = l_Lean_Syntax_node5(x_19, x_28, x_30, x_37, x_44, x_18, x_26); +x_46 = l_Lean_Syntax_node2(x_19, x_22, x_27, x_45); lean_inc(x_46); x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___boxed), 4, 1); lean_closure_set(x_47, 0, x_46); @@ -12988,17 +12984,17 @@ if (lean_obj_tag(x_55) == 0) lean_object* x_65; x_65 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_65); -x_18 = x_61; -x_19 = x_62; -x_20 = x_64; +x_18 = x_62; +x_19 = x_64; +x_20 = x_61; x_21 = lean_box(0); goto block_49; } else { -x_18 = x_61; -x_19 = x_62; -x_20 = x_64; +x_18 = x_62; +x_19 = x_64; +x_20 = x_61; x_21 = lean_box(0); goto block_49; } @@ -13679,19 +13675,19 @@ lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); lean_dec_ref(x_258); -lean_inc_ref(x_238); +lean_inc_ref(x_237); +lean_inc(x_245); lean_inc(x_248); -lean_inc(x_239); x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_239); -lean_ctor_set(x_260, 1, x_248); -lean_ctor_set(x_260, 2, x_238); +lean_ctor_set(x_260, 0, x_248); +lean_ctor_set(x_260, 1, x_245); +lean_ctor_set(x_260, 2, x_237); lean_inc_ref(x_260); -lean_inc(x_239); -x_261 = l_Lean_Syntax_node1(x_239, x_240, x_260); -lean_inc(x_239); -x_262 = l_Lean_Syntax_node2(x_239, x_241, x_246, x_260); -x_263 = l_Lean_Syntax_node2(x_239, x_245, x_261, x_262); +lean_inc(x_248); +x_261 = l_Lean_Syntax_node1(x_248, x_238, x_260); +lean_inc(x_248); +x_262 = l_Lean_Syntax_node2(x_248, x_243, x_242, x_260); +x_263 = l_Lean_Syntax_node2(x_248, x_236, x_261, x_262); x_264 = lean_unsigned_to_nat(2u); x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; x_266 = lean_array_push(x_265, x_247); @@ -13708,21 +13704,21 @@ lean_dec(x_10); x_272 = lean_ctor_get(x_271, 0); lean_inc(x_272); lean_dec_ref(x_271); -x_132 = x_269; -x_133 = x_236; -x_134 = x_237; -x_135 = x_238; -x_136 = x_253; -x_137 = x_242; -x_138 = x_264; -x_139 = x_243; -x_140 = x_244; -x_141 = x_259; +x_132 = x_237; +x_133 = x_239; +x_134 = x_264; +x_135 = x_240; +x_136 = x_241; +x_137 = x_269; +x_138 = x_244; +x_139 = x_259; +x_140 = x_245; +x_141 = x_253; x_142 = x_257; -x_143 = x_270; -x_144 = x_248; -x_145 = x_249; -x_146 = x_250; +x_143 = x_246; +x_144 = x_249; +x_145 = x_250; +x_146 = x_270; x_147 = x_272; x_148 = lean_box(0); goto block_229; @@ -13734,21 +13730,21 @@ lean_dec(x_10); x_273 = lean_ctor_get(x_234, 0); lean_inc(x_273); lean_dec_ref(x_234); -x_132 = x_269; -x_133 = x_236; -x_134 = x_237; -x_135 = x_238; -x_136 = x_253; -x_137 = x_242; -x_138 = x_264; -x_139 = x_243; -x_140 = x_244; -x_141 = x_259; +x_132 = x_237; +x_133 = x_239; +x_134 = x_264; +x_135 = x_240; +x_136 = x_241; +x_137 = x_269; +x_138 = x_244; +x_139 = x_259; +x_140 = x_245; +x_141 = x_253; x_142 = x_257; -x_143 = x_270; -x_144 = x_248; -x_145 = x_249; -x_146 = x_250; +x_143 = x_246; +x_144 = x_249; +x_145 = x_250; +x_146 = x_270; x_147 = x_273; x_148 = lean_box(0); goto block_229; @@ -13761,19 +13757,19 @@ lean_dec(x_257); lean_dec(x_255); lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec(x_248); lean_dec(x_247); lean_dec(x_246); lean_dec(x_245); lean_dec_ref(x_244); lean_dec(x_243); -lean_dec_ref(x_242); +lean_dec(x_242); lean_dec(x_241); -lean_dec(x_240); -lean_dec(x_239); -lean_dec_ref(x_238); -lean_dec(x_237); +lean_dec_ref(x_240); +lean_dec_ref(x_239); +lean_dec(x_238); +lean_dec_ref(x_237); lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); @@ -13803,19 +13799,19 @@ uint8_t x_277; lean_dec(x_255); lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec(x_248); lean_dec(x_247); lean_dec(x_246); lean_dec(x_245); lean_dec_ref(x_244); lean_dec(x_243); -lean_dec_ref(x_242); +lean_dec(x_242); lean_dec(x_241); -lean_dec(x_240); -lean_dec(x_239); -lean_dec_ref(x_238); -lean_dec(x_237); +lean_dec_ref(x_240); +lean_dec_ref(x_239); +lean_dec(x_238); +lean_dec_ref(x_237); lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); @@ -13845,19 +13841,19 @@ else uint8_t x_280; lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec(x_248); lean_dec(x_247); lean_dec(x_246); lean_dec(x_245); lean_dec_ref(x_244); lean_dec(x_243); -lean_dec_ref(x_242); +lean_dec(x_242); lean_dec(x_241); -lean_dec(x_240); -lean_dec(x_239); -lean_dec_ref(x_238); -lean_dec(x_237); +lean_dec_ref(x_240); +lean_dec_ref(x_239); +lean_dec(x_238); +lean_dec_ref(x_237); lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); @@ -13885,19 +13881,19 @@ return x_282; else { lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec(x_248); lean_dec(x_247); lean_dec(x_246); lean_dec(x_245); lean_dec_ref(x_244); lean_dec(x_243); -lean_dec_ref(x_242); +lean_dec(x_242); lean_dec(x_241); -lean_dec(x_240); -lean_dec(x_239); -lean_dec_ref(x_238); -lean_dec(x_237); +lean_dec_ref(x_240); +lean_dec_ref(x_239); +lean_dec(x_238); +lean_dec_ref(x_237); lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); @@ -13930,67 +13926,67 @@ x_299 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_300 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_301 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_302 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_289); +lean_inc(x_284); x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_289); +lean_ctor_set(x_303, 0, x_284); lean_ctor_set(x_303, 1, x_301); lean_ctor_set(x_303, 2, x_302); lean_inc_ref(x_303); -lean_inc(x_289); -x_304 = l_Lean_Syntax_node1(x_289, x_300, x_303); +lean_inc(x_284); +x_304 = l_Lean_Syntax_node1(x_284, x_300, x_303); x_305 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_306 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_307 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_308 = l_Lean_addMacroScope(x_291, x_307, x_288); -lean_inc(x_290); +x_308 = l_Lean_addMacroScope(x_291, x_307, x_285); lean_inc(x_289); +lean_inc(x_284); x_309 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_309, 0, x_289); +lean_ctor_set(x_309, 0, x_284); lean_ctor_set(x_309, 1, x_306); lean_ctor_set(x_309, 2, x_308); -lean_ctor_set(x_309, 3, x_290); -lean_inc(x_289); -x_310 = l_Lean_Syntax_node2(x_289, x_305, x_309, x_303); -x_311 = l_Lean_Syntax_node2(x_289, x_299, x_304, x_310); +lean_ctor_set(x_309, 3, x_289); +lean_inc(x_284); +x_310 = l_Lean_Syntax_node2(x_284, x_305, x_309, x_303); +x_311 = l_Lean_Syntax_node2(x_284, x_299, x_304, x_310); if (lean_obj_tag(x_234) == 0) { lean_object* x_312; x_312 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_312); -x_236 = x_285; -x_237 = x_286; -x_238 = x_302; -x_239 = x_294; -x_240 = x_300; -x_241 = x_305; -x_242 = x_296; -x_243 = x_287; -x_244 = x_297; -x_245 = x_299; -x_246 = x_284; +x_236 = x_299; +x_237 = x_302; +x_238 = x_300; +x_239 = x_296; +x_240 = x_297; +x_241 = x_287; +x_242 = x_288; +x_243 = x_305; +x_244 = x_298; +x_245 = x_301; +x_246 = x_286; x_247 = x_311; -x_248 = x_301; -x_249 = x_298; +x_248 = x_294; +x_249 = x_289; x_250 = x_290; x_251 = lean_box(0); goto block_283; } else { -x_236 = x_285; -x_237 = x_286; -x_238 = x_302; -x_239 = x_294; -x_240 = x_300; -x_241 = x_305; -x_242 = x_296; -x_243 = x_287; -x_244 = x_297; -x_245 = x_299; -x_246 = x_284; +x_236 = x_299; +x_237 = x_302; +x_238 = x_300; +x_239 = x_296; +x_240 = x_297; +x_241 = x_287; +x_242 = x_288; +x_243 = x_305; +x_244 = x_298; +x_245 = x_301; +x_246 = x_286; x_247 = x_311; -x_248 = x_301; -x_249 = x_298; +x_248 = x_294; +x_249 = x_289; x_250 = x_290; x_251 = lean_box(0); goto block_283; @@ -14131,13 +14127,13 @@ x_342 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lak x_343 = lean_ctor_get(x_342, 0); lean_inc(x_343); lean_dec_ref(x_342); -x_284 = x_337; -x_285 = x_329; -x_286 = x_331; -x_287 = x_324; -x_288 = x_341; -x_289 = x_339; -x_290 = x_328; +x_284 = x_339; +x_285 = x_341; +x_286 = x_324; +x_287 = x_331; +x_288 = x_337; +x_289 = x_328; +x_290 = x_329; x_291 = x_343; x_292 = lean_box(0); goto block_319; @@ -14150,13 +14146,13 @@ lean_inc(x_344); lean_dec_ref(x_340); x_345 = lean_ctor_get(x_234, 0); lean_inc(x_345); -x_284 = x_337; -x_285 = x_329; -x_286 = x_331; -x_287 = x_324; -x_288 = x_344; -x_289 = x_339; -x_290 = x_328; +x_284 = x_339; +x_285 = x_344; +x_286 = x_324; +x_287 = x_331; +x_288 = x_337; +x_289 = x_328; +x_290 = x_329; x_291 = x_345; x_292 = lean_box(0); goto block_319; @@ -14353,43 +14349,43 @@ return x_364; block_131: { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_28); -x_43 = l_Array_append___redArg(x_28, x_42); +lean_inc_ref(x_13); +x_43 = l_Array_append___redArg(x_13, x_42); lean_dec_ref(x_42); -lean_inc(x_23); -lean_inc(x_19); +lean_inc(x_37); +lean_inc(x_40); x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_19); -lean_ctor_set(x_44, 1, x_23); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_37); lean_ctor_set(x_44, 2, x_43); -lean_inc_n(x_22, 6); -lean_inc(x_36); +lean_inc_n(x_12, 6); lean_inc(x_19); -x_45 = l_Lean_Syntax_node7(x_19, x_36, x_44, x_22, x_22, x_22, x_22, x_22, x_22); +lean_inc(x_40); +x_45 = l_Lean_Syntax_node7(x_40, x_19, x_44, x_12, x_12, x_12, x_12, x_12, x_12); x_46 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; -lean_inc_ref(x_25); +lean_inc_ref(x_27); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_47 = l_Lean_Name_mkStr4(x_16, x_18, x_25, x_46); -lean_inc(x_19); +lean_inc_ref(x_32); +x_47 = l_Lean_Name_mkStr4(x_32, x_18, x_27, x_46); +lean_inc(x_40); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_19); +lean_ctor_set(x_48, 0, x_40); lean_ctor_set(x_48, 1, x_46); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; -lean_inc_ref(x_25); +lean_inc_ref(x_27); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_50 = l_Lean_Name_mkStr4(x_16, x_18, x_25, x_49); +lean_inc_ref(x_32); +x_50 = l_Lean_Name_mkStr4(x_32, x_18, x_27, x_49); x_51 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_52 = lean_box(2); -lean_inc(x_23); +lean_inc(x_37); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_23); +lean_ctor_set(x_53, 1, x_37); lean_ctor_set(x_53, 2, x_51); -x_54 = lean_mk_empty_array_with_capacity(x_17); -lean_inc(x_30); -x_55 = lean_array_push(x_54, x_30); +x_54 = lean_mk_empty_array_with_capacity(x_16); +lean_inc(x_41); +x_55 = lean_array_push(x_54, x_41); x_56 = lean_array_push(x_55, x_53); lean_inc(x_50); x_57 = lean_alloc_ctor(1, 3, 0); @@ -14397,194 +14393,194 @@ lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_50); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_25); +lean_inc_ref(x_27); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_59 = l_Lean_Name_mkStr4(x_16, x_18, x_25, x_58); +lean_inc_ref(x_32); +x_59 = l_Lean_Name_mkStr4(x_32, x_18, x_27, x_58); x_60 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_38); +lean_inc_ref(x_21); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_61 = l_Lean_Name_mkStr4(x_16, x_18, x_38, x_60); -lean_inc(x_15); +lean_inc_ref(x_32); +x_61 = l_Lean_Name_mkStr4(x_32, x_18, x_21, x_60); +lean_inc(x_31); lean_inc(x_61); -lean_inc(x_19); -x_62 = l_Lean_Syntax_node2(x_19, x_61, x_15, x_13); -lean_inc(x_23); -lean_inc(x_19); -x_63 = l_Lean_Syntax_node1(x_19, x_23, x_62); -lean_inc(x_22); +lean_inc(x_40); +x_62 = l_Lean_Syntax_node2(x_40, x_61, x_31, x_23); +lean_inc(x_37); +lean_inc(x_40); +x_63 = l_Lean_Syntax_node1(x_40, x_37, x_62); +lean_inc(x_12); lean_inc(x_59); -lean_inc(x_19); -x_64 = l_Lean_Syntax_node2(x_19, x_59, x_22, x_63); +lean_inc(x_40); +x_64 = l_Lean_Syntax_node2(x_40, x_59, x_12, x_63); x_65 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__60; -lean_inc_ref(x_25); +lean_inc_ref(x_27); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_66 = l_Lean_Name_mkStr4(x_16, x_18, x_25, x_65); +lean_inc_ref(x_32); +x_66 = l_Lean_Name_mkStr4(x_32, x_18, x_27, x_65); x_67 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__1; x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__2; -lean_inc_ref(x_24); -x_69 = l_Lean_Name_mkStr3(x_24, x_35, x_68); -lean_inc(x_31); +lean_inc_ref(x_17); +x_69 = l_Lean_Name_mkStr3(x_17, x_33, x_68); +lean_inc(x_22); lean_inc(x_69); -lean_inc(x_14); -x_70 = l_Lean_addMacroScope(x_14, x_69, x_31); +lean_inc(x_34); +x_70 = l_Lean_addMacroScope(x_34, x_69, x_22); lean_inc(x_39); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_69); lean_ctor_set(x_71, 1, x_39); -lean_inc(x_41); +lean_inc(x_24); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_41); -lean_inc(x_19); +lean_ctor_set(x_72, 1, x_24); +lean_inc(x_40); x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_19); +lean_ctor_set(x_73, 0, x_40); lean_ctor_set(x_73, 1, x_67); lean_ctor_set(x_73, 2, x_70); lean_ctor_set(x_73, 3, x_72); -lean_inc(x_23); -lean_inc(x_19); -x_74 = l_Lean_Syntax_node4(x_19, x_23, x_33, x_27, x_20, x_12); -lean_inc(x_19); -x_75 = l_Lean_Syntax_node2(x_19, x_34, x_73, x_74); +lean_inc(x_37); +lean_inc(x_40); +x_74 = l_Lean_Syntax_node4(x_40, x_37, x_28, x_35, x_29, x_25); +lean_inc(x_40); +x_75 = l_Lean_Syntax_node2(x_40, x_15, x_73, x_74); x_76 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__62; x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__63; lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_78 = l_Lean_Name_mkStr4(x_16, x_18, x_76, x_77); -lean_inc_n(x_22, 2); -lean_inc(x_19); -x_79 = l_Lean_Syntax_node2(x_19, x_78, x_22, x_22); -lean_inc(x_22); -lean_inc(x_79); +lean_inc_ref(x_32); +x_78 = l_Lean_Name_mkStr4(x_32, x_18, x_76, x_77); +lean_inc_n(x_12, 2); lean_inc(x_40); +x_79 = l_Lean_Syntax_node2(x_40, x_78, x_12, x_12); +lean_inc(x_12); +lean_inc(x_79); +lean_inc(x_38); lean_inc(x_66); -lean_inc(x_19); -x_80 = l_Lean_Syntax_node4(x_19, x_66, x_40, x_75, x_79, x_22); -lean_inc(x_19); -x_81 = l_Lean_Syntax_node4(x_19, x_47, x_48, x_57, x_64, x_80); -lean_inc(x_29); -lean_inc(x_19); -x_82 = l_Lean_Syntax_node2(x_19, x_29, x_45, x_81); +lean_inc(x_40); +x_80 = l_Lean_Syntax_node4(x_40, x_66, x_38, x_75, x_79, x_12); +lean_inc(x_40); +x_81 = l_Lean_Syntax_node4(x_40, x_47, x_48, x_57, x_64, x_80); +lean_inc(x_36); +lean_inc(x_40); +x_82 = l_Lean_Syntax_node2(x_40, x_36, x_45, x_81); x_83 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_38); +lean_inc_ref(x_21); lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_84 = l_Lean_Name_mkStr4(x_16, x_18, x_38, x_83); +lean_inc_ref(x_32); +x_84 = l_Lean_Name_mkStr4(x_32, x_18, x_21, x_83); x_85 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_19); +lean_inc(x_40); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_19); +lean_ctor_set(x_86, 0, x_40); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Syntax_SepArray_ofElems(x_37, x_26); -lean_dec_ref(x_26); -x_88 = l_Array_append___redArg(x_28, x_87); +x_87 = l_Lean_Syntax_SepArray_ofElems(x_30, x_20); +lean_dec_ref(x_20); +x_88 = l_Array_append___redArg(x_13, x_87); lean_dec_ref(x_87); -lean_inc(x_23); -lean_inc(x_19); +lean_inc(x_37); +lean_inc(x_40); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_19); -lean_ctor_set(x_89, 1, x_23); +lean_ctor_set(x_89, 0, x_40); +lean_ctor_set(x_89, 1, x_37); lean_ctor_set(x_89, 2, x_88); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_19); +lean_inc(x_40); x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_19); +lean_ctor_set(x_91, 0, x_40); lean_ctor_set(x_91, 1, x_90); -lean_inc(x_19); -x_92 = l_Lean_Syntax_node3(x_19, x_84, x_86, x_89, x_91); -lean_inc(x_23); -lean_inc(x_19); -x_93 = l_Lean_Syntax_node1(x_19, x_23, x_92); -lean_inc_n(x_22, 6); -lean_inc(x_19); -x_94 = l_Lean_Syntax_node7(x_19, x_36, x_22, x_93, x_22, x_22, x_22, x_22, x_22); +lean_inc(x_40); +x_92 = l_Lean_Syntax_node3(x_40, x_84, x_86, x_89, x_91); +lean_inc(x_37); +lean_inc(x_40); +x_93 = l_Lean_Syntax_node1(x_40, x_37, x_92); +lean_inc_n(x_12, 6); +lean_inc(x_40); +x_94 = l_Lean_Syntax_node7(x_40, x_19, x_12, x_93, x_12, x_12, x_12, x_12, x_12); x_95 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; lean_inc_ref(x_18); -lean_inc_ref(x_16); -x_96 = l_Lean_Name_mkStr4(x_16, x_18, x_25, x_95); +lean_inc_ref(x_32); +x_96 = l_Lean_Name_mkStr4(x_32, x_18, x_27, x_95); x_97 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_19); +lean_inc(x_40); x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_19); +lean_ctor_set(x_98, 0, x_40); lean_ctor_set(x_98, 1, x_97); x_99 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_100 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_31); -lean_inc(x_14); -x_101 = l_Lean_addMacroScope(x_14, x_100, x_31); -lean_inc(x_41); -lean_inc(x_19); +lean_inc(x_22); +lean_inc(x_34); +x_101 = l_Lean_addMacroScope(x_34, x_100, x_22); +lean_inc(x_24); +lean_inc(x_40); x_102 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_102, 0, x_19); +lean_ctor_set(x_102, 0, x_40); lean_ctor_set(x_102, 1, x_99); lean_ctor_set(x_102, 2, x_101); -lean_ctor_set(x_102, 3, x_41); -lean_inc(x_22); -lean_inc(x_19); -x_103 = l_Lean_Syntax_node2(x_19, x_50, x_102, x_22); +lean_ctor_set(x_102, 3, x_24); +lean_inc(x_12); +lean_inc(x_40); +x_103 = l_Lean_Syntax_node2(x_40, x_50, x_102, x_12); x_104 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__5; x_105 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_31); -lean_inc(x_14); -x_107 = l_Lean_addMacroScope(x_14, x_106, x_31); -x_108 = l_Lean_Name_mkStr2(x_24, x_104); +lean_inc(x_22); +lean_inc(x_34); +x_107 = l_Lean_addMacroScope(x_34, x_106, x_22); +x_108 = l_Lean_Name_mkStr2(x_17, x_104); lean_inc(x_108); x_109 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_109, 0, x_108); lean_ctor_set(x_109, 1, x_39); x_110 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_110, 0, x_108); -lean_inc(x_41); +lean_inc(x_24); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_41); +lean_ctor_set(x_111, 1, x_24); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_109); lean_ctor_set(x_112, 1, x_111); -lean_inc(x_19); +lean_inc(x_40); x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_19); +lean_ctor_set(x_113, 0, x_40); lean_ctor_set(x_113, 1, x_105); lean_ctor_set(x_113, 2, x_107); lean_ctor_set(x_113, 3, x_112); -lean_inc(x_19); -x_114 = l_Lean_Syntax_node2(x_19, x_61, x_15, x_113); -lean_inc(x_23); -lean_inc(x_19); -x_115 = l_Lean_Syntax_node1(x_19, x_23, x_114); -lean_inc(x_22); -lean_inc(x_19); -x_116 = l_Lean_Syntax_node2(x_19, x_59, x_22, x_115); +lean_inc(x_40); +x_114 = l_Lean_Syntax_node2(x_40, x_61, x_31, x_113); +lean_inc(x_37); +lean_inc(x_40); +x_115 = l_Lean_Syntax_node1(x_40, x_37, x_114); +lean_inc(x_12); +lean_inc(x_40); +x_116 = l_Lean_Syntax_node2(x_40, x_59, x_12, x_115); x_117 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_118 = l_Lean_Name_mkStr4(x_16, x_18, x_38, x_117); +x_118 = l_Lean_Name_mkStr4(x_32, x_18, x_21, x_117); x_119 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_19); +lean_inc(x_40); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_19); +lean_ctor_set(x_120, 0, x_40); lean_ctor_set(x_120, 1, x_119); x_121 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_123 = l_Lean_addMacroScope(x_14, x_122, x_31); -lean_inc(x_19); +x_123 = l_Lean_addMacroScope(x_34, x_122, x_22); +lean_inc(x_40); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_19); +lean_ctor_set(x_124, 0, x_40); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_41); -lean_inc(x_19); -x_125 = l_Lean_Syntax_node3(x_19, x_118, x_30, x_120, x_124); -lean_inc(x_22); -lean_inc(x_19); -x_126 = l_Lean_Syntax_node4(x_19, x_66, x_40, x_125, x_79, x_22); -lean_inc(x_19); -x_127 = l_Lean_Syntax_node5(x_19, x_96, x_98, x_103, x_116, x_126, x_22); -lean_inc(x_19); -x_128 = l_Lean_Syntax_node2(x_19, x_29, x_94, x_127); -x_129 = l_Lean_Syntax_node3(x_19, x_23, x_32, x_82, x_128); +lean_ctor_set(x_124, 3, x_24); +lean_inc(x_40); +x_125 = l_Lean_Syntax_node3(x_40, x_118, x_41, x_120, x_124); +lean_inc(x_12); +lean_inc(x_40); +x_126 = l_Lean_Syntax_node4(x_40, x_66, x_38, x_125, x_79, x_12); +lean_inc(x_40); +x_127 = l_Lean_Syntax_node5(x_40, x_96, x_98, x_103, x_116, x_126, x_12); +lean_inc(x_40); +x_128 = l_Lean_Syntax_node2(x_40, x_36, x_94, x_127); +x_129 = l_Lean_Syntax_node3(x_40, x_37, x_26, x_82, x_128); x_130 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_130, 0, x_129); return x_130; @@ -14594,13 +14590,13 @@ return x_130; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; x_149 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__0; x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; -lean_inc_ref(x_135); -lean_inc(x_144); +lean_inc_ref(x_132); +lean_inc(x_140); lean_inc(x_142); x_151 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_151, 0, x_142); -lean_ctor_set(x_151, 1, x_144); -lean_ctor_set(x_151, 2, x_135); +lean_ctor_set(x_151, 1, x_140); +lean_ctor_set(x_151, 2, x_132); x_152 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); @@ -14613,15 +14609,15 @@ lean_ctor_set(x_155, 0, x_142); lean_ctor_set(x_155, 1, x_154); x_156 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_157 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_141); +lean_inc(x_139); lean_inc(x_147); -x_158 = l_Lean_addMacroScope(x_147, x_157, x_141); +x_158 = l_Lean_addMacroScope(x_147, x_157, x_139); x_159 = lean_box(0); x_160 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__3; -lean_inc(x_146); +lean_inc(x_144); x_161 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_146); +lean_ctor_set(x_161, 1, x_144); lean_inc(x_142); x_162 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_162, 0, x_142); @@ -14629,15 +14625,15 @@ lean_ctor_set(x_162, 1, x_156); lean_ctor_set(x_162, 2, x_158); lean_ctor_set(x_162, 3, x_161); x_163 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_145); -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_164 = l_Lean_Name_mkStr4(x_137, x_140, x_145, x_163); +lean_inc_ref(x_138); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_164 = l_Lean_Name_mkStr4(x_133, x_135, x_138, x_163); x_165 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_145); -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_166 = l_Lean_Name_mkStr4(x_137, x_140, x_145, x_165); +lean_inc_ref(x_138); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_166 = l_Lean_Name_mkStr4(x_133, x_135, x_138, x_165); x_167 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; lean_inc(x_142); x_168 = lean_alloc_ctor(2, 2, 0); @@ -14646,39 +14642,39 @@ lean_ctor_set(x_168, 1, x_167); x_169 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_170 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_171 = lean_box(0); -lean_inc(x_141); +lean_inc(x_139); lean_inc(x_147); -x_172 = l_Lean_addMacroScope(x_147, x_171, x_141); +x_172 = l_Lean_addMacroScope(x_147, x_171, x_139); x_173 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__1; x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__59; -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_176 = l_Lean_Name_mkStr3(x_137, x_140, x_175); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_176 = l_Lean_Name_mkStr3(x_133, x_135, x_175); x_177 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_177, 0, x_176); x_178 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_137); -x_179 = l_Lean_Name_mkStr3(x_137, x_178, x_175); +lean_inc_ref(x_133); +x_179 = l_Lean_Name_mkStr3(x_133, x_178, x_175); x_180 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_180, 0, x_179); -lean_inc_ref(x_137); -x_181 = l_Lean_Name_mkStr2(x_137, x_178); +lean_inc_ref(x_133); +x_181 = l_Lean_Name_mkStr2(x_133, x_178); x_182 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_182, 0, x_181); -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_183 = l_Lean_Name_mkStr2(x_137, x_140); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_183 = l_Lean_Name_mkStr2(x_133, x_135); x_184 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_184, 0, x_183); -lean_inc_ref(x_137); -x_185 = l_Lean_Name_mkStr1(x_137); +lean_inc_ref(x_133); +x_185 = l_Lean_Name_mkStr1(x_133); x_186 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_186, 0, x_185); -lean_inc(x_146); +lean_inc(x_144); x_187 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_146); +lean_ctor_set(x_187, 1, x_144); x_188 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_188, 0, x_184); lean_ctor_set(x_188, 1, x_187); @@ -14717,14 +14713,14 @@ lean_inc(x_142); x_201 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_201, 0, x_142); lean_ctor_set(x_201, 1, x_200); -lean_inc(x_134); -lean_inc(x_144); +lean_inc(x_136); +lean_inc(x_140); lean_inc(x_142); -x_202 = l_Lean_Syntax_node1(x_142, x_144, x_134); +x_202 = l_Lean_Syntax_node1(x_142, x_140, x_136); lean_inc(x_199); -lean_inc(x_144); +lean_inc(x_140); lean_inc(x_142); -x_203 = l_Lean_Syntax_node3(x_142, x_144, x_199, x_201, x_202); +x_203 = l_Lean_Syntax_node3(x_142, x_140, x_199, x_201, x_202); x_204 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; lean_inc(x_142); x_205 = lean_alloc_ctor(2, 2, 0); @@ -14738,21 +14734,21 @@ x_208 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_208, 0, x_142); lean_ctor_set(x_208, 1, x_207); x_209 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_145); -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_210 = l_Lean_Name_mkStr4(x_137, x_140, x_145, x_209); +lean_inc_ref(x_138); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_210 = l_Lean_Name_mkStr4(x_133, x_135, x_138, x_209); x_211 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_212 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_141); +lean_inc(x_139); lean_inc(x_147); -x_213 = l_Lean_addMacroScope(x_147, x_212, x_141); +x_213 = l_Lean_addMacroScope(x_147, x_212, x_139); x_214 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__8; x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_146); +lean_inc(x_144); x_216 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_146); +lean_ctor_set(x_216, 1, x_144); x_217 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_217, 0, x_214); lean_ctor_set(x_217, 1, x_216); @@ -14762,61 +14758,61 @@ lean_ctor_set(x_218, 0, x_142); lean_ctor_set(x_218, 1, x_211); lean_ctor_set(x_218, 2, x_213); lean_ctor_set(x_218, 3, x_217); -lean_inc(x_143); -lean_inc(x_144); +lean_inc(x_146); +lean_inc(x_140); lean_inc(x_142); -x_219 = l_Lean_Syntax_node1(x_142, x_144, x_143); +x_219 = l_Lean_Syntax_node1(x_142, x_140, x_146); lean_inc(x_210); lean_inc(x_142); x_220 = l_Lean_Syntax_node2(x_142, x_210, x_218, x_219); lean_inc_ref(x_208); lean_inc_ref(x_155); -lean_inc(x_139); +lean_inc(x_143); lean_inc_ref(x_151); lean_inc(x_142); -x_221 = l_Lean_Syntax_node8(x_142, x_150, x_151, x_153, x_139, x_155, x_162, x_206, x_208, x_220); +x_221 = l_Lean_Syntax_node8(x_142, x_150, x_151, x_153, x_143, x_155, x_162, x_206, x_208, x_220); x_222 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_223 = l_Lean_Name_mkStr4(x_137, x_140, x_175, x_222); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_223 = l_Lean_Name_mkStr4(x_133, x_135, x_175, x_222); x_224 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_140); -lean_inc_ref(x_137); -x_225 = l_Lean_Name_mkStr4(x_137, x_140, x_175, x_224); +lean_inc_ref(x_135); +lean_inc_ref(x_133); +x_225 = l_Lean_Name_mkStr4(x_133, x_135, x_175, x_224); if (lean_obj_tag(x_5) == 0) { lean_object* x_226; x_226 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_12 = x_133; -x_13 = x_136; -x_14 = x_147; -x_15 = x_155; -x_16 = x_137; -x_17 = x_138; -x_18 = x_140; -x_19 = x_142; -x_20 = x_143; -x_21 = lean_box(0); -x_22 = x_151; -x_23 = x_144; -x_24 = x_149; -x_25 = x_175; -x_26 = x_132; -x_27 = x_134; -x_28 = x_135; -x_29 = x_223; -x_30 = x_139; -x_31 = x_141; -x_32 = x_221; -x_33 = x_199; -x_34 = x_210; -x_35 = x_173; -x_36 = x_225; -x_37 = x_200; -x_38 = x_145; +x_12 = x_151; +x_13 = x_132; +x_14 = lean_box(0); +x_15 = x_210; +x_16 = x_134; +x_17 = x_149; +x_18 = x_135; +x_19 = x_225; +x_20 = x_137; +x_21 = x_138; +x_22 = x_139; +x_23 = x_141; +x_24 = x_144; +x_25 = x_145; +x_26 = x_221; +x_27 = x_175; +x_28 = x_199; +x_29 = x_146; +x_30 = x_200; +x_31 = x_155; +x_32 = x_133; +x_33 = x_173; +x_34 = x_147; +x_35 = x_136; +x_36 = x_223; +x_37 = x_140; +x_38 = x_208; x_39 = x_159; -x_40 = x_208; -x_41 = x_146; +x_40 = x_142; +x_41 = x_143; x_42 = x_226; goto block_131; } @@ -14827,36 +14823,36 @@ x_227 = lean_ctor_get(x_5, 0); lean_inc(x_227); lean_dec_ref(x_5); x_228 = l_Array_mkArray1___redArg(x_227); -x_12 = x_133; -x_13 = x_136; -x_14 = x_147; -x_15 = x_155; -x_16 = x_137; -x_17 = x_138; -x_18 = x_140; -x_19 = x_142; -x_20 = x_143; -x_21 = lean_box(0); -x_22 = x_151; -x_23 = x_144; -x_24 = x_149; -x_25 = x_175; -x_26 = x_132; -x_27 = x_134; -x_28 = x_135; -x_29 = x_223; -x_30 = x_139; -x_31 = x_141; -x_32 = x_221; -x_33 = x_199; -x_34 = x_210; -x_35 = x_173; -x_36 = x_225; -x_37 = x_200; -x_38 = x_145; +x_12 = x_151; +x_13 = x_132; +x_14 = lean_box(0); +x_15 = x_210; +x_16 = x_134; +x_17 = x_149; +x_18 = x_135; +x_19 = x_225; +x_20 = x_137; +x_21 = x_138; +x_22 = x_139; +x_23 = x_141; +x_24 = x_144; +x_25 = x_145; +x_26 = x_221; +x_27 = x_175; +x_28 = x_199; +x_29 = x_146; +x_30 = x_200; +x_31 = x_155; +x_32 = x_133; +x_33 = x_173; +x_34 = x_147; +x_35 = x_136; +x_36 = x_223; +x_37 = x_140; +x_38 = x_208; x_39 = x_159; -x_40 = x_208; -x_41 = x_146; +x_40 = x_142; +x_41 = x_143; x_42 = x_228; goto block_131; } @@ -15434,524 +15430,266 @@ x_2 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand___regBuiltin__ return x_2; } } -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___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_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_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_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - lean_ctor_release(x_1, 2); - x_13 = x_1; -} else { - lean_dec_ref(x_1); - x_13 = lean_box(0); +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = l_Lake_LeanExeConfig_instConfigInfo; +x_9 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__1; +lean_inc(x_4); +x_10 = l_Lean_Syntax_isOfKind(x_4, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_12 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_11, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_12; } -lean_inc_ref(x_9); -x_14 = l___private_Lake_DSL_DeclUtil_0__Lake_DSL_mkConfigFields(x_2, x_12, x_7, x_9, x_10); -lean_dec(x_12); -if (lean_obj_tag(x_14) == 0) +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_50; lean_object* x_51; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -x_16 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__0; -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_16); -x_50 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__7; -if (lean_obj_tag(x_8) == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_4, x_13); +lean_inc(x_14); +x_15 = l_Lean_Syntax_matchesNull(x_14, x_13); +if (x_15 == 0) { -lean_object* x_73; -x_73 = lean_box(0); -x_51 = x_73; -goto block_72; +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(1u); +lean_inc(x_14); +x_17 = l_Lean_Syntax_matchesNull(x_14, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_14); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_19 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_18, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_19; } else { -uint8_t x_74; -x_74 = !lean_is_exclusive(x_8); -if (x_74 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Syntax_getArg(x_14, x_13); +lean_dec(x_14); +x_21 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__5; +lean_inc(x_20); +x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); +if (x_22 == 0) { -x_51 = x_8; -goto block_72; +lean_object* x_23; uint8_t x_24; +x_23 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__7; +lean_inc(x_20); +x_24 = l_Lean_Syntax_isOfKind(x_20, x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_20); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_25 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_26 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_25, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_26; } else { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_8, 0); -lean_inc(x_75); -lean_dec(x_8); -x_76 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_76, 0, x_75); -x_51 = x_76; -goto block_72; -} +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = l_Lean_Syntax_getArg(x_20, x_13); +x_28 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__9; +lean_inc(x_27); +x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_27); +lean_dec(x_20); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_30 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_31 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_30, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_31; } -block_49: +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; 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 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__1; -x_23 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__2; -x_24 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; -x_25 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_19); -if (lean_is_scalar(x_13)) { - x_26 = lean_alloc_ctor(1, 3, 0); -} else { - x_26 = x_13; - lean_ctor_set_tag(x_26, 1); +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = l_Lean_Syntax_getArg(x_27, x_16); +x_33 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__11; +lean_inc(x_32); +x_34 = l_Lean_Syntax_isOfKind(x_32, x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_32); +lean_dec(x_27); +lean_dec(x_20); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_35 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_36 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_35, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_36; } -lean_ctor_set(x_26, 0, x_19); -lean_ctor_set(x_26, 1, x_24); -lean_ctor_set(x_26, 2, x_25); -lean_inc_ref_n(x_26, 7); -lean_inc(x_19); -x_27 = l_Lean_Syntax_node7(x_19, x_23, x_26, x_26, x_26, x_26, x_26, x_26, x_26); -x_28 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__3; -x_29 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_19); -x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_19); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__4; -x_32 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; -lean_inc(x_18); -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_18); -lean_ctor_set(x_33, 1, x_24); -lean_ctor_set(x_33, 2, x_32); -x_34 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_35 = lean_array_push(x_34, x_3); -x_36 = lean_array_push(x_35, x_33); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_18); -lean_ctor_set(x_37, 1, x_31); -lean_ctor_set(x_37, 2, x_36); -x_38 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__5; -x_39 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__58; -x_40 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_19); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_19); -lean_ctor_set(x_41, 1, x_40); -lean_inc(x_19); -x_42 = l_Lean_Syntax_node2(x_19, x_39, x_41, x_4); -lean_inc(x_19); -x_43 = l_Lean_Syntax_node1(x_19, x_24, x_42); -lean_inc_ref(x_26); -lean_inc(x_19); -x_44 = l_Lean_Syntax_node2(x_19, x_38, x_26, x_43); -lean_inc(x_19); -x_45 = l_Lean_Syntax_node5(x_19, x_28, x_30, x_37, x_44, x_20, x_26); -x_46 = l_Lean_Syntax_node2(x_19, x_22, x_27, x_45); -lean_inc(x_46); -x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___boxed), 4, 1); -lean_closure_set(x_47, 0, x_46); -x_48 = l_Lean_Elab_Command_withMacroExpansion___redArg(x_5, x_46, x_47, x_9, x_10); -return x_48; -} -block_72: -{ -lean_object* x_52; -x_52 = l_Lean_Elab_Command_getRef___redArg(x_9); -if (lean_obj_tag(x_52) == 0) +else { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -lean_dec_ref(x_52); -x_54 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_9); -if (lean_obj_tag(x_54) == 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_48; uint8_t x_49; +x_37 = l_Lean_Syntax_getArg(x_27, x_13); +lean_dec(x_27); +x_38 = l_Lean_Syntax_getArg(x_32, x_13); +lean_dec(x_32); +x_48 = l_Lean_Syntax_getArg(x_20, x_16); +lean_dec(x_20); +x_49 = l_Lean_Syntax_isNone(x_48); +if (x_49 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; -lean_dec_ref(x_54); -x_55 = lean_ctor_get(x_9, 5); -x_56 = l_Lean_mkOptionalNode(x_51); -x_57 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__8; -x_58 = lean_array_push(x_57, x_17); -x_59 = lean_array_push(x_58, x_15); -x_60 = lean_array_push(x_59, x_56); -x_61 = lean_box(2); -x_62 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_50); -lean_ctor_set(x_62, 2, x_60); -x_63 = 0; -x_64 = l_Lean_SourceInfo_fromRef(x_53, x_63); -lean_dec(x_53); -if (lean_obj_tag(x_55) == 0) +uint8_t x_50; +lean_inc(x_48); +x_50 = l_Lean_Syntax_matchesNull(x_48, x_16); +if (x_50 == 0) { -lean_object* x_65; -x_65 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); -lean_dec_ref(x_65); -x_18 = x_61; -x_19 = x_64; -x_20 = x_62; -x_21 = lean_box(0); -goto block_49; +lean_object* x_51; lean_object* x_52; +lean_dec(x_48); +lean_dec(x_38); +lean_dec(x_37); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_52 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_51, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_52; } else { -x_18 = x_61; -x_19 = x_64; -x_20 = x_62; -x_21 = lean_box(0); -goto block_49; -} -} -else +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Syntax_getArg(x_48, x_13); +lean_dec(x_48); +x_54 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__66; +lean_inc(x_53); +x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); +if (x_55 == 0) { -uint8_t x_66; +lean_object* x_56; lean_object* x_57; lean_dec(x_53); -lean_dec(x_51); -lean_dec_ref(x_17); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_38); +lean_dec(x_37); lean_dec(x_3); -x_66 = !lean_is_exclusive(x_54); -if (x_66 == 0) -{ -return x_54; +lean_dec(x_2); +lean_dec(x_1); +x_56 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_57 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_56, x_5, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_57; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_54, 0); -lean_inc(x_67); -lean_dec(x_54); -x_68 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_68, 0, x_67); -return x_68; +lean_object* x_58; +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_53); +x_39 = x_58; +x_40 = x_5; +x_41 = x_6; +x_42 = lean_box(0); +goto block_47; } } } else { -uint8_t x_69; -lean_dec(x_51); -lean_dec_ref(x_17); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_69 = !lean_is_exclusive(x_52); -if (x_69 == 0) -{ -return x_52; +lean_object* x_59; +lean_dec(x_48); +x_59 = lean_box(0); +x_39 = x_59; +x_40 = x_5; +x_41 = x_6; +x_42 = lean_box(0); +goto block_47; } -else +block_47: { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_52, 0); -lean_inc(x_70); -lean_dec(x_52); -x_71 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = l_Lean_Syntax_getArgs(x_38); +lean_dec(x_38); +x_44 = l_Lean_Syntax_getHeadInfo(x_37); +lean_dec(x_37); +x_45 = l_Lean_Syntax_TSepArray_getElems___redArg(x_43); +lean_dec_ref(x_43); +x_46 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_44, x_45, x_39, x_40, x_41); +lean_dec_ref(x_45); +return x_46; +} } } } } else { -uint8_t x_77; -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = l_Lean_Syntax_getArg(x_20, x_16); +x_61 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__11; +lean_inc(x_60); +x_62 = l_Lean_Syntax_isOfKind(x_60, x_61); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_60); +lean_dec(x_20); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_63 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_64 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_63, x_5, x_6); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_77 = !lean_is_exclusive(x_14); -if (x_77 == 0) -{ -return x_14; +return x_64; } else { -lean_object* x_78; lean_object* x_79; -x_78 = lean_ctor_get(x_14, 0); -lean_inc(x_78); -lean_dec(x_14); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -return x_79; -} -} -} -} -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_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_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_65 = l_Lean_Syntax_getArg(x_20, x_13); +x_66 = l_Lean_Syntax_getArg(x_60, x_13); +lean_dec(x_60); +x_76 = lean_unsigned_to_nat(2u); +x_77 = l_Lean_Syntax_getArg(x_20, x_76); +lean_dec(x_20); +x_78 = l_Lean_Syntax_isNone(x_77); +if (x_78 == 0) { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = l_Lake_LeanExeConfig_instConfigInfo; -x_9 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__1; -lean_inc(x_4); -x_10 = l_Lean_Syntax_isOfKind(x_4, x_9); -if (x_10 == 0) +uint8_t x_79; +lean_inc(x_77); +x_79 = l_Lean_Syntax_matchesNull(x_77, x_16); +if (x_79 == 0) { -lean_object* x_11; lean_object* x_12; +lean_object* x_80; lean_object* x_81; +lean_dec(x_77); +lean_dec(x_66); +lean_dec(x_65); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_11 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_12 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_11, x_5, x_6); +x_80 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; +x_81 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_80, x_5, x_6); lean_dec(x_6); lean_dec(x_4); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_getArg(x_4, x_13); -lean_inc(x_14); -x_15 = l_Lean_Syntax_matchesNull(x_14, x_13); -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(1u); -lean_inc(x_14); -x_17 = l_Lean_Syntax_matchesNull(x_14, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_14); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_18 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_19 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_18, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Syntax_getArg(x_14, x_13); -lean_dec(x_14); -x_21 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__5; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__7; -lean_inc(x_20); -x_24 = l_Lean_Syntax_isOfKind(x_20, x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -lean_dec(x_20); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_26 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_25, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_26; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_20, x_13); -x_28 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__9; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_27); -lean_dec(x_20); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_30 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_31 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_30, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = l_Lean_Syntax_getArg(x_27, x_16); -x_33 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__11; -lean_inc(x_32); -x_34 = l_Lean_Syntax_isOfKind(x_32, x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_20); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_35 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_36 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_35, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -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_48; uint8_t x_49; -x_37 = l_Lean_Syntax_getArg(x_27, x_13); -lean_dec(x_27); -x_38 = l_Lean_Syntax_getArg(x_32, x_13); -lean_dec(x_32); -x_48 = l_Lean_Syntax_getArg(x_20, x_16); -lean_dec(x_20); -x_49 = l_Lean_Syntax_isNone(x_48); -if (x_49 == 0) -{ -uint8_t x_50; -lean_inc(x_48); -x_50 = l_Lean_Syntax_matchesNull(x_48, x_16); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; -lean_dec(x_48); -lean_dec(x_38); -lean_dec(x_37); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_51 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_52 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_51, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_52; -} -else -{ -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = l_Lean_Syntax_getArg(x_48, x_13); -lean_dec(x_48); -x_54 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__66; -lean_inc(x_53); -x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; -lean_dec(x_53); -lean_dec(x_38); -lean_dec(x_37); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_57 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_56, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_57; -} -else -{ -lean_object* x_58; -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_53); -x_39 = x_58; -x_40 = x_5; -x_41 = x_6; -x_42 = lean_box(0); -goto block_47; -} -} -} -else -{ -lean_object* x_59; -lean_dec(x_48); -x_59 = lean_box(0); -x_39 = x_59; -x_40 = x_5; -x_41 = x_6; -x_42 = lean_box(0); -goto block_47; -} -block_47: -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = l_Lean_Syntax_getArgs(x_38); -lean_dec(x_38); -x_44 = l_Lean_Syntax_getHeadInfo(x_37); -lean_dec(x_37); -x_45 = l_Lean_Syntax_TSepArray_getElems___redArg(x_43); -lean_dec_ref(x_43); -x_46 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_44, x_45, x_39, x_40, x_41); -lean_dec_ref(x_45); -return x_46; -} -} -} -} -} -else -{ -lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_60 = l_Lean_Syntax_getArg(x_20, x_16); -x_61 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__11; -lean_inc(x_60); -x_62 = l_Lean_Syntax_isOfKind(x_60, x_61); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; -lean_dec(x_60); -lean_dec(x_20); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_63 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_64 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_63, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_64; -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_65 = l_Lean_Syntax_getArg(x_20, x_13); -x_66 = l_Lean_Syntax_getArg(x_60, x_13); -lean_dec(x_60); -x_76 = lean_unsigned_to_nat(2u); -x_77 = l_Lean_Syntax_getArg(x_20, x_76); -lean_dec(x_20); -x_78 = l_Lean_Syntax_isNone(x_77); -if (x_78 == 0) -{ -uint8_t x_79; -lean_inc(x_77); -x_79 = l_Lean_Syntax_matchesNull(x_77, x_16); -if (x_79 == 0) -{ -lean_object* x_80; lean_object* x_81; -lean_dec(x_77); -lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_80 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___closed__3; -x_81 = l_Lean_throwErrorAt___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__0___redArg(x_4, x_80, x_5, x_6); -lean_dec(x_6); -lean_dec(x_4); -return x_81; +return x_81; } else { @@ -16009,7 +15747,7 @@ x_72 = l_Lean_Syntax_getHeadInfo(x_65); lean_dec(x_65); x_73 = l_Lean_Syntax_TSepArray_getElems___redArg(x_71); lean_dec_ref(x_71); -x_74 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_72, x_73, x_67, x_68, x_69); +x_74 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_72, x_73, x_67, x_68, x_69); lean_dec_ref(x_73); return x_74; } @@ -16024,7 +15762,7 @@ lean_dec(x_14); x_89 = lean_box(2); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_91 = lean_box(0); -x_92 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_89, x_90, x_91, x_5, x_6); +x_92 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_89, x_90, x_91, x_5, x_6); return x_92; } } @@ -16105,22 +15843,22 @@ lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); lean_dec_ref(x_258); -lean_inc_ref(x_236); +lean_inc_ref(x_250); lean_inc(x_249); -lean_inc(x_240); +lean_inc(x_244); x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_240); +lean_ctor_set(x_260, 0, x_244); lean_ctor_set(x_260, 1, x_249); -lean_ctor_set(x_260, 2, x_236); +lean_ctor_set(x_260, 2, x_250); lean_inc_ref(x_260); -lean_inc(x_240); -x_261 = l_Lean_Syntax_node1(x_240, x_241, x_260); -lean_inc(x_240); -x_262 = l_Lean_Syntax_node2(x_240, x_238, x_244, x_260); -x_263 = l_Lean_Syntax_node2(x_240, x_245, x_261, x_262); +lean_inc(x_244); +x_261 = l_Lean_Syntax_node1(x_244, x_239, x_260); +lean_inc(x_244); +x_262 = l_Lean_Syntax_node2(x_244, x_243, x_238, x_260); +x_263 = l_Lean_Syntax_node2(x_244, x_240, x_261, x_262); x_264 = lean_unsigned_to_nat(2u); x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_266 = lean_array_push(x_265, x_243); +x_266 = lean_array_push(x_265, x_245); x_267 = lean_array_push(x_266, x_263); x_268 = l_Lake_DSL_expandAttrs(x_6); x_269 = l_Array_append___redArg(x_267, x_268); @@ -16134,25 +15872,25 @@ lean_dec(x_10); x_272 = lean_ctor_get(x_271, 0); lean_inc(x_272); lean_dec_ref(x_271); -x_132 = x_236; -x_133 = x_237; -x_134 = x_259; -x_135 = x_239; -x_136 = x_270; -x_137 = x_253; -x_138 = x_264; -x_139 = x_242; -x_140 = x_246; -x_141 = x_247; -x_142 = x_248; -x_143 = x_269; -x_144 = x_257; -x_145 = x_249; -x_146 = x_250; -x_147 = x_272; -x_148 = lean_box(0); -goto block_229; -} +x_132 = x_253; +x_133 = x_236; +x_134 = x_237; +x_135 = x_241; +x_136 = x_242; +x_137 = x_259; +x_138 = x_270; +x_139 = x_257; +x_140 = x_269; +x_141 = x_246; +x_142 = x_264; +x_143 = x_247; +x_144 = x_249; +x_145 = x_250; +x_146 = x_248; +x_147 = x_272; +x_148 = lean_box(0); +goto block_229; +} else { lean_object* x_273; @@ -16160,21 +15898,21 @@ lean_dec(x_10); x_273 = lean_ctor_get(x_234, 0); lean_inc(x_273); lean_dec_ref(x_234); -x_132 = x_236; -x_133 = x_237; -x_134 = x_259; -x_135 = x_239; -x_136 = x_270; -x_137 = x_253; -x_138 = x_264; -x_139 = x_242; -x_140 = x_246; -x_141 = x_247; -x_142 = x_248; -x_143 = x_269; -x_144 = x_257; -x_145 = x_249; -x_146 = x_250; +x_132 = x_253; +x_133 = x_236; +x_134 = x_237; +x_135 = x_241; +x_136 = x_242; +x_137 = x_259; +x_138 = x_270; +x_139 = x_257; +x_140 = x_269; +x_141 = x_246; +x_142 = x_264; +x_143 = x_247; +x_144 = x_249; +x_145 = x_250; +x_146 = x_248; x_147 = x_273; x_148 = lean_box(0); goto block_229; @@ -16186,11 +15924,11 @@ uint8_t x_274; lean_dec(x_257); lean_dec(x_255); lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); +lean_dec_ref(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); @@ -16199,8 +15937,8 @@ lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec_ref(x_236); +lean_dec(x_237); +lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec(x_6); @@ -16228,11 +15966,11 @@ else uint8_t x_277; lean_dec(x_255); lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); +lean_dec_ref(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); @@ -16241,8 +15979,8 @@ lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec_ref(x_236); +lean_dec(x_237); +lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -16270,11 +16008,11 @@ else { uint8_t x_280; lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); +lean_dec_ref(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); @@ -16283,8 +16021,8 @@ lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec_ref(x_236); +lean_dec(x_237); +lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -16310,11 +16048,11 @@ return x_282; } else { -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); +lean_dec_ref(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); @@ -16323,8 +16061,8 @@ lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec_ref(x_236); +lean_dec(x_237); +lean_dec(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -16356,68 +16094,68 @@ x_299 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_300 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_301 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_302 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_289); +lean_inc(x_287); x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_289); +lean_ctor_set(x_303, 0, x_287); lean_ctor_set(x_303, 1, x_301); lean_ctor_set(x_303, 2, x_302); lean_inc_ref(x_303); -lean_inc(x_289); -x_304 = l_Lean_Syntax_node1(x_289, x_300, x_303); +lean_inc(x_287); +x_304 = l_Lean_Syntax_node1(x_287, x_300, x_303); x_305 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_306 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_307 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_308 = l_Lean_addMacroScope(x_291, x_307, x_288); -lean_inc(x_286); -lean_inc(x_289); +x_308 = l_Lean_addMacroScope(x_291, x_307, x_284); +lean_inc(x_290); +lean_inc(x_287); x_309 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_309, 0, x_289); +lean_ctor_set(x_309, 0, x_287); lean_ctor_set(x_309, 1, x_306); lean_ctor_set(x_309, 2, x_308); -lean_ctor_set(x_309, 3, x_286); -lean_inc(x_289); -x_310 = l_Lean_Syntax_node2(x_289, x_305, x_309, x_303); -x_311 = l_Lean_Syntax_node2(x_289, x_299, x_304, x_310); +lean_ctor_set(x_309, 3, x_290); +lean_inc(x_287); +x_310 = l_Lean_Syntax_node2(x_287, x_305, x_309, x_303); +x_311 = l_Lean_Syntax_node2(x_287, x_299, x_304, x_310); if (lean_obj_tag(x_234) == 0) { lean_object* x_312; x_312 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_312); -x_236 = x_302; -x_237 = x_296; -x_238 = x_305; -x_239 = x_285; -x_240 = x_294; -x_241 = x_300; -x_242 = x_298; -x_243 = x_311; -x_244 = x_284; -x_245 = x_299; -x_246 = x_286; -x_247 = x_297; -x_248 = x_287; +x_236 = x_286; +x_237 = x_285; +x_238 = x_288; +x_239 = x_300; +x_240 = x_299; +x_241 = x_289; +x_242 = x_296; +x_243 = x_305; +x_244 = x_294; +x_245 = x_311; +x_246 = x_298; +x_247 = x_290; +x_248 = x_297; x_249 = x_301; -x_250 = x_290; +x_250 = x_302; x_251 = lean_box(0); goto block_283; } else { -x_236 = x_302; -x_237 = x_296; -x_238 = x_305; -x_239 = x_285; -x_240 = x_294; -x_241 = x_300; -x_242 = x_298; -x_243 = x_311; -x_244 = x_284; -x_245 = x_299; -x_246 = x_286; -x_247 = x_297; -x_248 = x_287; +x_236 = x_286; +x_237 = x_285; +x_238 = x_288; +x_239 = x_300; +x_240 = x_299; +x_241 = x_289; +x_242 = x_296; +x_243 = x_305; +x_244 = x_294; +x_245 = x_311; +x_246 = x_298; +x_247 = x_290; +x_248 = x_297; x_249 = x_301; -x_250 = x_290; +x_250 = x_302; x_251 = lean_box(0); goto block_283; } @@ -16557,13 +16295,13 @@ x_342 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lak x_343 = lean_ctor_get(x_342, 0); lean_inc(x_343); lean_dec_ref(x_342); -x_284 = x_337; -x_285 = x_331; -x_286 = x_328; -x_287 = x_329; -x_288 = x_341; -x_289 = x_339; -x_290 = x_324; +x_284 = x_341; +x_285 = x_324; +x_286 = x_329; +x_287 = x_339; +x_288 = x_337; +x_289 = x_331; +x_290 = x_328; x_291 = x_343; x_292 = lean_box(0); goto block_319; @@ -16576,13 +16314,13 @@ lean_inc(x_344); lean_dec_ref(x_340); x_345 = lean_ctor_get(x_234, 0); lean_inc(x_345); -x_284 = x_337; -x_285 = x_331; -x_286 = x_328; -x_287 = x_329; -x_288 = x_344; -x_289 = x_339; -x_290 = x_324; +x_284 = x_344; +x_285 = x_324; +x_286 = x_329; +x_287 = x_339; +x_288 = x_337; +x_289 = x_331; +x_290 = x_328; x_291 = x_345; x_292 = lean_box(0); goto block_319; @@ -16779,43 +16517,43 @@ return x_364; block_131: { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_26); -x_43 = l_Array_append___redArg(x_26, x_42); +lean_inc_ref(x_41); +x_43 = l_Array_append___redArg(x_41, x_42); lean_dec_ref(x_42); -lean_inc(x_24); -lean_inc(x_22); +lean_inc(x_40); +lean_inc(x_18); x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_22); -lean_ctor_set(x_44, 1, x_24); +lean_ctor_set(x_44, 0, x_18); +lean_ctor_set(x_44, 1, x_40); lean_ctor_set(x_44, 2, x_43); -lean_inc_n(x_38, 6); -lean_inc(x_23); -lean_inc(x_22); -x_45 = l_Lean_Syntax_node7(x_22, x_23, x_44, x_38, x_38, x_38, x_38, x_38, x_38); +lean_inc_n(x_14, 6); +lean_inc(x_36); +lean_inc(x_18); +x_45 = l_Lean_Syntax_node7(x_18, x_36, x_44, x_14, x_14, x_14, x_14, x_14, x_14); x_46 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; +lean_inc_ref(x_25); +lean_inc_ref(x_23); lean_inc_ref(x_31); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_47 = l_Lean_Name_mkStr4(x_12, x_18, x_31, x_46); -lean_inc(x_22); +x_47 = l_Lean_Name_mkStr4(x_31, x_23, x_25, x_46); +lean_inc(x_18); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_22); +lean_ctor_set(x_48, 0, x_18); lean_ctor_set(x_48, 1, x_46); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; +lean_inc_ref(x_25); +lean_inc_ref(x_23); lean_inc_ref(x_31); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_50 = l_Lean_Name_mkStr4(x_12, x_18, x_31, x_49); +x_50 = l_Lean_Name_mkStr4(x_31, x_23, x_25, x_49); x_51 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_52 = lean_box(2); -lean_inc(x_24); +lean_inc(x_40); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_24); +lean_ctor_set(x_53, 1, x_40); lean_ctor_set(x_53, 2, x_51); -x_54 = lean_mk_empty_array_with_capacity(x_15); -lean_inc(x_25); -x_55 = lean_array_push(x_54, x_25); +x_54 = lean_mk_empty_array_with_capacity(x_21); +lean_inc(x_13); +x_55 = lean_array_push(x_54, x_13); x_56 = lean_array_push(x_55, x_53); lean_inc(x_50); x_57 = lean_alloc_ctor(1, 3, 0); @@ -16823,194 +16561,194 @@ lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_50); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; +lean_inc_ref(x_25); +lean_inc_ref(x_23); lean_inc_ref(x_31); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_59 = l_Lean_Name_mkStr4(x_12, x_18, x_31, x_58); +x_59 = l_Lean_Name_mkStr4(x_31, x_23, x_25, x_58); x_60 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_36); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_61 = l_Lean_Name_mkStr4(x_12, x_18, x_36, x_60); -lean_inc(x_35); +lean_inc_ref(x_38); +lean_inc_ref(x_23); +lean_inc_ref(x_31); +x_61 = l_Lean_Name_mkStr4(x_31, x_23, x_38, x_60); +lean_inc(x_33); lean_inc(x_61); -lean_inc(x_22); -x_62 = l_Lean_Syntax_node2(x_22, x_61, x_35, x_30); -lean_inc(x_24); -lean_inc(x_22); -x_63 = l_Lean_Syntax_node1(x_22, x_24, x_62); -lean_inc(x_38); +lean_inc(x_18); +x_62 = l_Lean_Syntax_node2(x_18, x_61, x_33, x_29); +lean_inc(x_40); +lean_inc(x_18); +x_63 = l_Lean_Syntax_node1(x_18, x_40, x_62); +lean_inc(x_14); lean_inc(x_59); -lean_inc(x_22); -x_64 = l_Lean_Syntax_node2(x_22, x_59, x_38, x_63); +lean_inc(x_18); +x_64 = l_Lean_Syntax_node2(x_18, x_59, x_14, x_63); x_65 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__60; +lean_inc_ref(x_25); +lean_inc_ref(x_23); lean_inc_ref(x_31); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_66 = l_Lean_Name_mkStr4(x_12, x_18, x_31, x_65); +x_66 = l_Lean_Name_mkStr4(x_31, x_23, x_25, x_65); x_67 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__1; x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__2; -lean_inc_ref(x_20); -x_69 = l_Lean_Name_mkStr3(x_20, x_29, x_68); -lean_inc(x_13); +lean_inc_ref(x_39); +x_69 = l_Lean_Name_mkStr3(x_39, x_20, x_68); +lean_inc(x_17); lean_inc(x_69); -lean_inc(x_37); -x_70 = l_Lean_addMacroScope(x_37, x_69, x_13); -lean_inc(x_14); +lean_inc(x_27); +x_70 = l_Lean_addMacroScope(x_27, x_69, x_17); +lean_inc(x_35); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_14); -lean_inc(x_16); +lean_ctor_set(x_71, 1, x_35); +lean_inc(x_22); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_16); -lean_inc(x_22); +lean_ctor_set(x_72, 1, x_22); +lean_inc(x_18); x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_22); +lean_ctor_set(x_73, 0, x_18); lean_ctor_set(x_73, 1, x_67); lean_ctor_set(x_73, 2, x_70); lean_ctor_set(x_73, 3, x_72); -lean_inc(x_24); -lean_inc(x_22); -x_74 = l_Lean_Syntax_node4(x_22, x_24, x_21, x_27, x_32, x_40); -lean_inc(x_22); -x_75 = l_Lean_Syntax_node2(x_22, x_28, x_73, x_74); +lean_inc(x_40); +lean_inc(x_18); +x_74 = l_Lean_Syntax_node4(x_18, x_40, x_16, x_15, x_34, x_30); +lean_inc(x_18); +x_75 = l_Lean_Syntax_node2(x_18, x_26, x_73, x_74); x_76 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__62; x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__63; -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_78 = l_Lean_Name_mkStr4(x_12, x_18, x_76, x_77); -lean_inc_n(x_38, 2); -lean_inc(x_22); -x_79 = l_Lean_Syntax_node2(x_22, x_78, x_38, x_38); -lean_inc(x_38); +lean_inc_ref(x_23); +lean_inc_ref(x_31); +x_78 = l_Lean_Name_mkStr4(x_31, x_23, x_76, x_77); +lean_inc_n(x_14, 2); +lean_inc(x_18); +x_79 = l_Lean_Syntax_node2(x_18, x_78, x_14, x_14); +lean_inc(x_14); lean_inc(x_79); -lean_inc(x_17); +lean_inc(x_24); lean_inc(x_66); -lean_inc(x_22); -x_80 = l_Lean_Syntax_node4(x_22, x_66, x_17, x_75, x_79, x_38); -lean_inc(x_22); -x_81 = l_Lean_Syntax_node4(x_22, x_47, x_48, x_57, x_64, x_80); -lean_inc(x_34); -lean_inc(x_22); -x_82 = l_Lean_Syntax_node2(x_22, x_34, x_45, x_81); +lean_inc(x_18); +x_80 = l_Lean_Syntax_node4(x_18, x_66, x_24, x_75, x_79, x_14); +lean_inc(x_18); +x_81 = l_Lean_Syntax_node4(x_18, x_47, x_48, x_57, x_64, x_80); +lean_inc(x_12); +lean_inc(x_18); +x_82 = l_Lean_Syntax_node2(x_18, x_12, x_45, x_81); x_83 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_36); -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_84 = l_Lean_Name_mkStr4(x_12, x_18, x_36, x_83); +lean_inc_ref(x_38); +lean_inc_ref(x_23); +lean_inc_ref(x_31); +x_84 = l_Lean_Name_mkStr4(x_31, x_23, x_38, x_83); x_85 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_22); +lean_inc(x_18); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_22); +lean_ctor_set(x_86, 0, x_18); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Syntax_SepArray_ofElems(x_33, x_41); -lean_dec_ref(x_41); -x_88 = l_Array_append___redArg(x_26, x_87); +x_87 = l_Lean_Syntax_SepArray_ofElems(x_32, x_37); +lean_dec_ref(x_37); +x_88 = l_Array_append___redArg(x_41, x_87); lean_dec_ref(x_87); -lean_inc(x_24); -lean_inc(x_22); +lean_inc(x_40); +lean_inc(x_18); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_22); -lean_ctor_set(x_89, 1, x_24); +lean_ctor_set(x_89, 0, x_18); +lean_ctor_set(x_89, 1, x_40); lean_ctor_set(x_89, 2, x_88); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_22); +lean_inc(x_18); x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_22); +lean_ctor_set(x_91, 0, x_18); lean_ctor_set(x_91, 1, x_90); -lean_inc(x_22); -x_92 = l_Lean_Syntax_node3(x_22, x_84, x_86, x_89, x_91); -lean_inc(x_24); -lean_inc(x_22); -x_93 = l_Lean_Syntax_node1(x_22, x_24, x_92); -lean_inc_n(x_38, 6); -lean_inc(x_22); -x_94 = l_Lean_Syntax_node7(x_22, x_23, x_38, x_93, x_38, x_38, x_38, x_38, x_38); +lean_inc(x_18); +x_92 = l_Lean_Syntax_node3(x_18, x_84, x_86, x_89, x_91); +lean_inc(x_40); +lean_inc(x_18); +x_93 = l_Lean_Syntax_node1(x_18, x_40, x_92); +lean_inc_n(x_14, 6); +lean_inc(x_18); +x_94 = l_Lean_Syntax_node7(x_18, x_36, x_14, x_93, x_14, x_14, x_14, x_14, x_14); x_95 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; -lean_inc_ref(x_18); -lean_inc_ref(x_12); -x_96 = l_Lean_Name_mkStr4(x_12, x_18, x_31, x_95); +lean_inc_ref(x_23); +lean_inc_ref(x_31); +x_96 = l_Lean_Name_mkStr4(x_31, x_23, x_25, x_95); x_97 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_22); +lean_inc(x_18); x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_22); +lean_ctor_set(x_98, 0, x_18); lean_ctor_set(x_98, 1, x_97); x_99 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_100 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_13); -lean_inc(x_37); -x_101 = l_Lean_addMacroScope(x_37, x_100, x_13); -lean_inc(x_16); +lean_inc(x_17); +lean_inc(x_27); +x_101 = l_Lean_addMacroScope(x_27, x_100, x_17); lean_inc(x_22); +lean_inc(x_18); x_102 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_102, 0, x_22); +lean_ctor_set(x_102, 0, x_18); lean_ctor_set(x_102, 1, x_99); lean_ctor_set(x_102, 2, x_101); -lean_ctor_set(x_102, 3, x_16); -lean_inc(x_38); -lean_inc(x_22); -x_103 = l_Lean_Syntax_node2(x_22, x_50, x_102, x_38); +lean_ctor_set(x_102, 3, x_22); +lean_inc(x_14); +lean_inc(x_18); +x_103 = l_Lean_Syntax_node2(x_18, x_50, x_102, x_14); x_104 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__5; x_105 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_13); -lean_inc(x_37); -x_107 = l_Lean_addMacroScope(x_37, x_106, x_13); -x_108 = l_Lean_Name_mkStr2(x_20, x_104); +lean_inc(x_17); +lean_inc(x_27); +x_107 = l_Lean_addMacroScope(x_27, x_106, x_17); +x_108 = l_Lean_Name_mkStr2(x_39, x_104); lean_inc(x_108); x_109 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_14); +lean_ctor_set(x_109, 1, x_35); x_110 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_110, 0, x_108); -lean_inc(x_16); +lean_inc(x_22); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_16); +lean_ctor_set(x_111, 1, x_22); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_109); lean_ctor_set(x_112, 1, x_111); -lean_inc(x_22); +lean_inc(x_18); x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_22); +lean_ctor_set(x_113, 0, x_18); lean_ctor_set(x_113, 1, x_105); lean_ctor_set(x_113, 2, x_107); lean_ctor_set(x_113, 3, x_112); -lean_inc(x_22); -x_114 = l_Lean_Syntax_node2(x_22, x_61, x_35, x_113); -lean_inc(x_24); -lean_inc(x_22); -x_115 = l_Lean_Syntax_node1(x_22, x_24, x_114); -lean_inc(x_38); -lean_inc(x_22); -x_116 = l_Lean_Syntax_node2(x_22, x_59, x_38, x_115); +lean_inc(x_18); +x_114 = l_Lean_Syntax_node2(x_18, x_61, x_33, x_113); +lean_inc(x_40); +lean_inc(x_18); +x_115 = l_Lean_Syntax_node1(x_18, x_40, x_114); +lean_inc(x_14); +lean_inc(x_18); +x_116 = l_Lean_Syntax_node2(x_18, x_59, x_14, x_115); x_117 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_118 = l_Lean_Name_mkStr4(x_12, x_18, x_36, x_117); +x_118 = l_Lean_Name_mkStr4(x_31, x_23, x_38, x_117); x_119 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_22); +lean_inc(x_18); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_22); +lean_ctor_set(x_120, 0, x_18); lean_ctor_set(x_120, 1, x_119); x_121 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_123 = l_Lean_addMacroScope(x_37, x_122, x_13); -lean_inc(x_22); +x_123 = l_Lean_addMacroScope(x_27, x_122, x_17); +lean_inc(x_18); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_22); +lean_ctor_set(x_124, 0, x_18); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_16); -lean_inc(x_22); -x_125 = l_Lean_Syntax_node3(x_22, x_118, x_25, x_120, x_124); -lean_inc(x_38); -lean_inc(x_22); -x_126 = l_Lean_Syntax_node4(x_22, x_66, x_17, x_125, x_79, x_38); -lean_inc(x_22); -x_127 = l_Lean_Syntax_node5(x_22, x_96, x_98, x_103, x_116, x_126, x_38); -lean_inc(x_22); -x_128 = l_Lean_Syntax_node2(x_22, x_34, x_94, x_127); -x_129 = l_Lean_Syntax_node3(x_22, x_24, x_19, x_82, x_128); +lean_ctor_set(x_124, 3, x_22); +lean_inc(x_18); +x_125 = l_Lean_Syntax_node3(x_18, x_118, x_13, x_120, x_124); +lean_inc(x_14); +lean_inc(x_18); +x_126 = l_Lean_Syntax_node4(x_18, x_66, x_24, x_125, x_79, x_14); +lean_inc(x_18); +x_127 = l_Lean_Syntax_node5(x_18, x_96, x_98, x_103, x_116, x_126, x_14); +lean_inc(x_18); +x_128 = l_Lean_Syntax_node2(x_18, x_12, x_94, x_127); +x_129 = l_Lean_Syntax_node3(x_18, x_40, x_28, x_82, x_128); x_130 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_130, 0, x_129); return x_130; @@ -17020,91 +16758,91 @@ return x_130; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; x_149 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__0; x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; -lean_inc_ref(x_132); -lean_inc(x_145); +lean_inc_ref(x_145); lean_inc(x_144); +lean_inc(x_139); x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_144); -lean_ctor_set(x_151, 1, x_145); -lean_ctor_set(x_151, 2, x_132); +lean_ctor_set(x_151, 0, x_139); +lean_ctor_set(x_151, 1, x_144); +lean_ctor_set(x_151, 2, x_145); x_152 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; -lean_inc(x_144); +lean_inc(x_139); x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_144); +lean_ctor_set(x_153, 0, x_139); lean_ctor_set(x_153, 1, x_152); x_154 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_144); +lean_inc(x_139); x_155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_155, 0, x_144); +lean_ctor_set(x_155, 0, x_139); lean_ctor_set(x_155, 1, x_154); x_156 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_157 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_134); +lean_inc(x_137); lean_inc(x_147); -x_158 = l_Lean_addMacroScope(x_147, x_157, x_134); +x_158 = l_Lean_addMacroScope(x_147, x_157, x_137); x_159 = lean_box(0); x_160 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__3; -lean_inc(x_140); +lean_inc(x_143); x_161 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_140); -lean_inc(x_144); +lean_ctor_set(x_161, 1, x_143); +lean_inc(x_139); x_162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_162, 0, x_144); +lean_ctor_set(x_162, 0, x_139); lean_ctor_set(x_162, 1, x_156); lean_ctor_set(x_162, 2, x_158); lean_ctor_set(x_162, 3, x_161); x_163 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_139); lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_164 = l_Lean_Name_mkStr4(x_133, x_141, x_139, x_163); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_164 = l_Lean_Name_mkStr4(x_136, x_146, x_141, x_163); x_165 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_139); lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_166 = l_Lean_Name_mkStr4(x_133, x_141, x_139, x_165); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_166 = l_Lean_Name_mkStr4(x_136, x_146, x_141, x_165); x_167 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_144); +lean_inc(x_139); x_168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_168, 0, x_144); +lean_ctor_set(x_168, 0, x_139); lean_ctor_set(x_168, 1, x_167); x_169 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_170 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_171 = lean_box(0); -lean_inc(x_134); +lean_inc(x_137); lean_inc(x_147); -x_172 = l_Lean_addMacroScope(x_147, x_171, x_134); +x_172 = l_Lean_addMacroScope(x_147, x_171, x_137); x_173 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__1; x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__59; -lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_176 = l_Lean_Name_mkStr3(x_133, x_141, x_175); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_176 = l_Lean_Name_mkStr3(x_136, x_146, x_175); x_177 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_177, 0, x_176); x_178 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_133); -x_179 = l_Lean_Name_mkStr3(x_133, x_178, x_175); +lean_inc_ref(x_136); +x_179 = l_Lean_Name_mkStr3(x_136, x_178, x_175); x_180 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_180, 0, x_179); -lean_inc_ref(x_133); -x_181 = l_Lean_Name_mkStr2(x_133, x_178); +lean_inc_ref(x_136); +x_181 = l_Lean_Name_mkStr2(x_136, x_178); x_182 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_182, 0, x_181); -lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_183 = l_Lean_Name_mkStr2(x_133, x_141); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_183 = l_Lean_Name_mkStr2(x_136, x_146); x_184 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_184, 0, x_183); -lean_inc_ref(x_133); -x_185 = l_Lean_Name_mkStr1(x_133); +lean_inc_ref(x_136); +x_185 = l_Lean_Name_mkStr1(x_136); x_186 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_186, 0, x_185); -lean_inc(x_140); +lean_inc(x_143); x_187 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_140); +lean_ctor_set(x_187, 1, x_143); x_188 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_188, 0, x_184); lean_ctor_set(x_188, 1, x_187); @@ -17120,129 +16858,129 @@ lean_ctor_set(x_191, 1, x_190); x_192 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_192, 0, x_174); lean_ctor_set(x_192, 1, x_191); -lean_inc(x_144); +lean_inc(x_139); x_193 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_193, 0, x_144); +lean_ctor_set(x_193, 0, x_139); lean_ctor_set(x_193, 1, x_170); lean_ctor_set(x_193, 2, x_172); lean_ctor_set(x_193, 3, x_192); -lean_inc(x_144); -x_194 = l_Lean_Syntax_node1(x_144, x_169, x_193); -lean_inc(x_144); -x_195 = l_Lean_Syntax_node2(x_144, x_166, x_168, x_194); +lean_inc(x_139); +x_194 = l_Lean_Syntax_node1(x_139, x_169, x_193); +lean_inc(x_139); +x_195 = l_Lean_Syntax_node2(x_139, x_166, x_168, x_194); x_196 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__37; x_197 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__38; -lean_inc(x_144); +lean_inc(x_139); x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_144); +lean_ctor_set(x_198, 0, x_139); lean_ctor_set(x_198, 1, x_197); -lean_inc(x_144); -x_199 = l_Lean_Syntax_node1(x_144, x_196, x_198); +lean_inc(x_139); +x_199 = l_Lean_Syntax_node1(x_139, x_196, x_198); x_200 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -lean_inc(x_144); +lean_inc(x_139); x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_144); +lean_ctor_set(x_201, 0, x_139); lean_ctor_set(x_201, 1, x_200); lean_inc(x_135); -lean_inc(x_145); lean_inc(x_144); -x_202 = l_Lean_Syntax_node1(x_144, x_145, x_135); +lean_inc(x_139); +x_202 = l_Lean_Syntax_node1(x_139, x_144, x_135); lean_inc(x_199); -lean_inc(x_145); lean_inc(x_144); -x_203 = l_Lean_Syntax_node3(x_144, x_145, x_199, x_201, x_202); +lean_inc(x_139); +x_203 = l_Lean_Syntax_node3(x_139, x_144, x_199, x_201, x_202); x_204 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_144); +lean_inc(x_139); x_205 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_205, 0, x_144); +lean_ctor_set(x_205, 0, x_139); lean_ctor_set(x_205, 1, x_204); -lean_inc(x_144); -x_206 = l_Lean_Syntax_node3(x_144, x_164, x_195, x_203, x_205); +lean_inc(x_139); +x_206 = l_Lean_Syntax_node3(x_139, x_164, x_195, x_203, x_205); x_207 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_144); +lean_inc(x_139); x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_144); +lean_ctor_set(x_208, 0, x_139); lean_ctor_set(x_208, 1, x_207); x_209 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_139); lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_210 = l_Lean_Name_mkStr4(x_133, x_141, x_139, x_209); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_210 = l_Lean_Name_mkStr4(x_136, x_146, x_141, x_209); x_211 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_212 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_134); +lean_inc(x_137); lean_inc(x_147); -x_213 = l_Lean_addMacroScope(x_147, x_212, x_134); +x_213 = l_Lean_addMacroScope(x_147, x_212, x_137); x_214 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__8; x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_140); +lean_inc(x_143); x_216 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_140); +lean_ctor_set(x_216, 1, x_143); x_217 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_217, 0, x_214); lean_ctor_set(x_217, 1, x_216); -lean_inc(x_144); +lean_inc(x_139); x_218 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_218, 0, x_144); +lean_ctor_set(x_218, 0, x_139); lean_ctor_set(x_218, 1, x_211); lean_ctor_set(x_218, 2, x_213); lean_ctor_set(x_218, 3, x_217); -lean_inc(x_136); -lean_inc(x_145); +lean_inc(x_138); lean_inc(x_144); -x_219 = l_Lean_Syntax_node1(x_144, x_145, x_136); +lean_inc(x_139); +x_219 = l_Lean_Syntax_node1(x_139, x_144, x_138); lean_inc(x_210); -lean_inc(x_144); -x_220 = l_Lean_Syntax_node2(x_144, x_210, x_218, x_219); +lean_inc(x_139); +x_220 = l_Lean_Syntax_node2(x_139, x_210, x_218, x_219); lean_inc_ref(x_208); lean_inc_ref(x_155); -lean_inc(x_146); +lean_inc(x_134); lean_inc_ref(x_151); -lean_inc(x_144); -x_221 = l_Lean_Syntax_node8(x_144, x_150, x_151, x_153, x_146, x_155, x_162, x_206, x_208, x_220); +lean_inc(x_139); +x_221 = l_Lean_Syntax_node8(x_139, x_150, x_151, x_153, x_134, x_155, x_162, x_206, x_208, x_220); x_222 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_223 = l_Lean_Name_mkStr4(x_133, x_141, x_175, x_222); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_223 = l_Lean_Name_mkStr4(x_136, x_146, x_175, x_222); x_224 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_141); -lean_inc_ref(x_133); -x_225 = l_Lean_Name_mkStr4(x_133, x_141, x_175, x_224); +lean_inc_ref(x_146); +lean_inc_ref(x_136); +x_225 = l_Lean_Name_mkStr4(x_136, x_146, x_175, x_224); if (lean_obj_tag(x_5) == 0) { lean_object* x_226; x_226 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_12 = x_133; +x_12 = x_223; x_13 = x_134; -x_14 = x_159; -x_15 = x_138; -x_16 = x_140; -x_17 = x_208; -x_18 = x_141; -x_19 = x_221; -x_20 = x_149; -x_21 = x_199; -x_22 = x_144; -x_23 = x_225; -x_24 = x_145; -x_25 = x_146; -x_26 = x_132; -x_27 = x_135; -x_28 = x_210; -x_29 = x_173; -x_30 = x_137; -x_31 = x_175; -x_32 = x_136; -x_33 = x_200; -x_34 = x_223; -x_35 = x_155; -x_36 = x_139; -x_37 = x_147; -x_38 = x_151; -x_39 = lean_box(0); -x_40 = x_142; -x_41 = x_143; +x_14 = x_151; +x_15 = x_135; +x_16 = x_199; +x_17 = x_137; +x_18 = x_139; +x_19 = lean_box(0); +x_20 = x_173; +x_21 = x_142; +x_22 = x_143; +x_23 = x_146; +x_24 = x_208; +x_25 = x_175; +x_26 = x_210; +x_27 = x_147; +x_28 = x_221; +x_29 = x_132; +x_30 = x_133; +x_31 = x_136; +x_32 = x_200; +x_33 = x_155; +x_34 = x_138; +x_35 = x_159; +x_36 = x_225; +x_37 = x_140; +x_38 = x_141; +x_39 = x_149; +x_40 = x_144; +x_41 = x_145; x_42 = x_226; goto block_131; } @@ -17253,36 +16991,36 @@ x_227 = lean_ctor_get(x_5, 0); lean_inc(x_227); lean_dec_ref(x_5); x_228 = l_Array_mkArray1___redArg(x_227); -x_12 = x_133; +x_12 = x_223; x_13 = x_134; -x_14 = x_159; -x_15 = x_138; -x_16 = x_140; -x_17 = x_208; -x_18 = x_141; -x_19 = x_221; -x_20 = x_149; -x_21 = x_199; -x_22 = x_144; -x_23 = x_225; -x_24 = x_145; -x_25 = x_146; -x_26 = x_132; -x_27 = x_135; -x_28 = x_210; -x_29 = x_173; -x_30 = x_137; -x_31 = x_175; -x_32 = x_136; -x_33 = x_200; -x_34 = x_223; -x_35 = x_155; -x_36 = x_139; -x_37 = x_147; -x_38 = x_151; -x_39 = lean_box(0); -x_40 = x_142; -x_41 = x_143; +x_14 = x_151; +x_15 = x_135; +x_16 = x_199; +x_17 = x_137; +x_18 = x_139; +x_19 = lean_box(0); +x_20 = x_173; +x_21 = x_142; +x_22 = x_143; +x_23 = x_146; +x_24 = x_208; +x_25 = x_175; +x_26 = x_210; +x_27 = x_147; +x_28 = x_221; +x_29 = x_132; +x_30 = x_133; +x_31 = x_136; +x_32 = x_200; +x_33 = x_155; +x_34 = x_138; +x_35 = x_159; +x_36 = x_225; +x_37 = x_140; +x_38 = x_141; +x_39 = x_149; +x_40 = x_144; +x_41 = x_145; x_42 = x_228; goto block_131; } @@ -17427,7 +17165,7 @@ lean_dec(x_14); lean_ctor_set(x_2, 7, x_29); lean_inc(x_3); lean_inc_ref(x_2); -x_30 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0(x_27, x_28, x_28, x_26, x_21, x_20, x_19, x_18, x_2, x_3); +x_30 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0(x_27, x_28, x_28, x_26, x_21, x_19, x_20, x_18, x_2, x_3); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -17506,7 +17244,7 @@ lean_ctor_set(x_51, 9, x_45); lean_ctor_set_uint8(x_51, sizeof(void*)*10, x_46); lean_inc(x_3); lean_inc_ref(x_51); -x_52 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0(x_48, x_49, x_49, x_47, x_21, x_20, x_19, x_18, x_51, x_3); +x_52 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0(x_48, x_49, x_49, x_47, x_21, x_19, x_20, x_18, x_51, x_3); if (lean_obj_tag(x_52) == 0) { lean_object* x_53; lean_object* x_54; lean_object* x_55; @@ -17581,8 +17319,8 @@ if (lean_obj_tag(x_65) == 0) { lean_object* x_66; x_66 = lean_box(0); -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_66; goto block_62; } @@ -17592,8 +17330,8 @@ uint8_t x_67; x_67 = !lean_is_exclusive(x_65); if (x_67 == 0) { -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_65; goto block_62; } @@ -17605,8 +17343,8 @@ lean_inc(x_68); lean_dec(x_65); x_69 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_69, 0, x_68); -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_69; goto block_62; } @@ -17652,15 +17390,6 @@ goto block_70; } } } -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___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_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_spec__0_spec__0___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_7); -return x_12; -} -} LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand_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) { _start: { @@ -17703,283 +17432,25 @@ x_3 = l_Lean_Name_str___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1() { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand__1___closed__0; -x_3 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___closed__1; -x_4 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1___closed__1; -x_5 = lean_alloc_closure((void*)(l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___boxed), 4, 0); -x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___redArg(x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1(); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___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_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - lean_ctor_release(x_1, 2); - x_13 = x_1; -} else { - lean_dec_ref(x_1); - x_13 = lean_box(0); -} -lean_inc_ref(x_9); -x_14 = l___private_Lake_DSL_DeclUtil_0__Lake_DSL_mkConfigFields(x_2, x_12, x_7, x_9, x_10); -lean_dec(x_12); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_50; lean_object* x_51; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec_ref(x_14); -x_16 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__0; -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_16); -x_50 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__7; -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_73; -x_73 = lean_box(0); -x_51 = x_73; -goto block_72; -} -else -{ -uint8_t x_74; -x_74 = !lean_is_exclusive(x_8); -if (x_74 == 0) -{ -x_51 = x_8; -goto block_72; -} -else -{ -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_8, 0); -lean_inc(x_75); -lean_dec(x_8); -x_76 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_76, 0, x_75); -x_51 = x_76; -goto block_72; -} -} -block_49: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; 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 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__1; -x_23 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__2; -x_24 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; -x_25 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_18); -if (lean_is_scalar(x_13)) { - x_26 = lean_alloc_ctor(1, 3, 0); -} else { - x_26 = x_13; - lean_ctor_set_tag(x_26, 1); -} -lean_ctor_set(x_26, 0, x_18); -lean_ctor_set(x_26, 1, x_24); -lean_ctor_set(x_26, 2, x_25); -lean_inc_ref_n(x_26, 7); -lean_inc(x_18); -x_27 = l_Lean_Syntax_node7(x_18, x_23, x_26, x_26, x_26, x_26, x_26, x_26, x_26); -x_28 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__3; -x_29 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_18); -x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_18); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__4; -x_32 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; -lean_inc(x_20); -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_24); -lean_ctor_set(x_33, 2, x_32); -x_34 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_35 = lean_array_push(x_34, x_3); -x_36 = lean_array_push(x_35, x_33); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_20); -lean_ctor_set(x_37, 1, x_31); -lean_ctor_set(x_37, 2, x_36); -x_38 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__5; -x_39 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__58; -x_40 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_18); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_18); -lean_ctor_set(x_41, 1, x_40); -lean_inc(x_18); -x_42 = l_Lean_Syntax_node2(x_18, x_39, x_41, x_4); -lean_inc(x_18); -x_43 = l_Lean_Syntax_node1(x_18, x_24, x_42); -lean_inc_ref(x_26); -lean_inc(x_18); -x_44 = l_Lean_Syntax_node2(x_18, x_38, x_26, x_43); -lean_inc(x_18); -x_45 = l_Lean_Syntax_node5(x_18, x_28, x_30, x_37, x_44, x_19, x_26); -x_46 = l_Lean_Syntax_node2(x_18, x_22, x_27, x_45); -lean_inc(x_46); -x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___boxed), 4, 1); -lean_closure_set(x_47, 0, x_46); -x_48 = l_Lean_Elab_Command_withMacroExpansion___redArg(x_5, x_46, x_47, x_9, x_10); -return x_48; -} -block_72: -{ -lean_object* x_52; -x_52 = l_Lean_Elab_Command_getRef___redArg(x_9); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -lean_dec_ref(x_52); -x_54 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_9); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; -lean_dec_ref(x_54); -x_55 = lean_ctor_get(x_9, 5); -x_56 = l_Lean_mkOptionalNode(x_51); -x_57 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__8; -x_58 = lean_array_push(x_57, x_17); -x_59 = lean_array_push(x_58, x_15); -x_60 = lean_array_push(x_59, x_56); -x_61 = lean_box(2); -x_62 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_50); -lean_ctor_set(x_62, 2, x_60); -x_63 = 0; -x_64 = l_Lean_SourceInfo_fromRef(x_53, x_63); -lean_dec(x_53); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_65; -x_65 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); -lean_dec_ref(x_65); -x_18 = x_64; -x_19 = x_62; -x_20 = x_61; -x_21 = lean_box(0); -goto block_49; -} -else -{ -x_18 = x_64; -x_19 = x_62; -x_20 = x_61; -x_21 = lean_box(0); -goto block_49; -} -} -else -{ -uint8_t x_66; -lean_dec(x_53); -lean_dec(x_51); -lean_dec_ref(x_17); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_66 = !lean_is_exclusive(x_54); -if (x_66 == 0) -{ -return x_54; -} -else -{ -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_54, 0); -lean_inc(x_67); -lean_dec(x_54); -x_68 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_68, 0, x_67); -return x_68; -} -} -} -else -{ -uint8_t x_69; -lean_dec(x_51); -lean_dec_ref(x_17); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_69 = !lean_is_exclusive(x_52); -if (x_69 == 0) -{ -return x_52; -} -else -{ -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_52, 0); -lean_inc(x_70); -lean_dec(x_52); -x_71 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; -} -} -} -} -else -{ -uint8_t x_77; -lean_dec(x_13); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_77 = !lean_is_exclusive(x_14); -if (x_77 == 0) -{ -return x_14; -} -else -{ -lean_object* x_78; lean_object* x_79; -x_78 = lean_ctor_get(x_14, 0); -lean_inc(x_78); -lean_dec(x_14); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -return x_79; -} +LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1() { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand__1___closed__0; +x_3 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___closed__1; +x_4 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1___closed__1; +x_5 = lean_alloc_closure((void*)(l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___boxed), 4, 0); +x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___redArg(x_2, x_3, x_4, x_5); +return x_6; } } +LEAN_EXPORT lean_object* l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand___regBuiltin___private_Lake_DSL_Targets_0__Lake_DSL_elabLeanExeCommand__1(); +return x_2; +} } LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_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: @@ -18183,7 +17654,7 @@ x_44 = l_Lean_Syntax_getHeadInfo(x_37); lean_dec(x_37); x_45 = l_Lean_Syntax_TSepArray_getElems___redArg(x_43); lean_dec_ref(x_43); -x_46 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_44, x_45, x_39, x_40, x_41); +x_46 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_44, x_45, x_39, x_40, x_41); lean_dec_ref(x_45); return x_46; } @@ -18298,7 +17769,7 @@ x_72 = l_Lean_Syntax_getHeadInfo(x_65); lean_dec(x_65); x_73 = l_Lean_Syntax_TSepArray_getElems___redArg(x_71); lean_dec_ref(x_71); -x_74 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_72, x_73, x_67, x_68, x_69); +x_74 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_72, x_73, x_67, x_68, x_69); lean_dec_ref(x_73); return x_74; } @@ -18313,7 +17784,7 @@ lean_dec(x_14); x_89 = lean_box(2); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_91 = lean_box(0); -x_92 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___lam__0(x_8, x_1, x_2, x_3, x_4, x_89, x_90, x_91, x_5, x_6); +x_92 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0(x_8, x_1, x_2, x_3, x_4, x_89, x_90, x_91, x_5, x_6); return x_92; } } @@ -18394,22 +17865,22 @@ lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); lean_dec_ref(x_258); -lean_inc_ref(x_243); -lean_inc(x_249); -lean_inc(x_242); +lean_inc_ref(x_236); +lean_inc(x_245); +lean_inc(x_239); x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_242); -lean_ctor_set(x_260, 1, x_249); -lean_ctor_set(x_260, 2, x_243); +lean_ctor_set(x_260, 0, x_239); +lean_ctor_set(x_260, 1, x_245); +lean_ctor_set(x_260, 2, x_236); lean_inc_ref(x_260); -lean_inc(x_242); -x_261 = l_Lean_Syntax_node1(x_242, x_239, x_260); -lean_inc(x_242); -x_262 = l_Lean_Syntax_node2(x_242, x_246, x_240, x_260); -x_263 = l_Lean_Syntax_node2(x_242, x_237, x_261, x_262); +lean_inc(x_239); +x_261 = l_Lean_Syntax_node1(x_239, x_241, x_260); +lean_inc(x_239); +x_262 = l_Lean_Syntax_node2(x_239, x_238, x_248, x_260); +x_263 = l_Lean_Syntax_node2(x_239, x_243, x_261, x_262); x_264 = lean_unsigned_to_nat(2u); x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_266 = lean_array_push(x_265, x_250); +x_266 = lean_array_push(x_265, x_242); x_267 = lean_array_push(x_266, x_263); x_268 = l_Lake_DSL_expandAttrs(x_6); x_269 = l_Array_append___redArg(x_267, x_268); @@ -18423,21 +17894,21 @@ lean_dec(x_10); x_272 = lean_ctor_get(x_271, 0); lean_inc(x_272); lean_dec_ref(x_271); -x_132 = x_236; -x_133 = x_238; -x_134 = x_269; -x_135 = x_270; -x_136 = x_253; -x_137 = x_241; -x_138 = x_243; -x_139 = x_244; -x_140 = x_247; +x_132 = x_257; +x_133 = x_236; +x_134 = x_259; +x_135 = x_237; +x_136 = x_270; +x_137 = x_269; +x_138 = x_240; +x_139 = x_253; +x_140 = x_244; x_141 = x_245; -x_142 = x_257; -x_143 = x_264; -x_144 = x_249; -x_145 = x_248; -x_146 = x_259; +x_142 = x_246; +x_143 = x_247; +x_144 = x_264; +x_145 = x_249; +x_146 = x_250; x_147 = x_272; x_148 = lean_box(0); goto block_229; @@ -18449,21 +17920,21 @@ lean_dec(x_10); x_273 = lean_ctor_get(x_234, 0); lean_inc(x_273); lean_dec_ref(x_234); -x_132 = x_236; -x_133 = x_238; -x_134 = x_269; -x_135 = x_270; -x_136 = x_253; -x_137 = x_241; -x_138 = x_243; -x_139 = x_244; -x_140 = x_247; +x_132 = x_257; +x_133 = x_236; +x_134 = x_259; +x_135 = x_237; +x_136 = x_270; +x_137 = x_269; +x_138 = x_240; +x_139 = x_253; +x_140 = x_244; x_141 = x_245; -x_142 = x_257; -x_143 = x_264; -x_144 = x_249; -x_145 = x_248; -x_146 = x_259; +x_142 = x_246; +x_143 = x_247; +x_144 = x_264; +x_145 = x_249; +x_146 = x_250; x_147 = x_273; x_148 = lean_box(0); goto block_229; @@ -18475,21 +17946,21 @@ uint8_t x_274; lean_dec(x_257); lean_dec(x_255); lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec_ref(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); -lean_dec_ref(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); +lean_dec(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); +lean_dec(x_245); +lean_dec_ref(x_244); +lean_dec(x_243); lean_dec(x_242); lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); lean_dec(x_237); -lean_dec(x_236); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec(x_6); @@ -18517,21 +17988,21 @@ else uint8_t x_277; lean_dec(x_255); lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec_ref(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); -lean_dec_ref(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); +lean_dec(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); +lean_dec(x_245); +lean_dec_ref(x_244); +lean_dec(x_243); lean_dec(x_242); lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); lean_dec(x_237); -lean_dec(x_236); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -18559,21 +18030,21 @@ else { uint8_t x_280; lean_dec(x_253); -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec_ref(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); -lean_dec_ref(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); +lean_dec(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); +lean_dec(x_245); +lean_dec_ref(x_244); +lean_dec(x_243); lean_dec(x_242); lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); lean_dec(x_237); -lean_dec(x_236); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -18599,21 +18070,21 @@ return x_282; } else { -lean_dec(x_250); +lean_dec_ref(x_250); lean_dec(x_249); -lean_dec_ref(x_248); -lean_dec_ref(x_247); -lean_dec(x_246); -lean_dec_ref(x_245); -lean_dec(x_244); -lean_dec_ref(x_243); +lean_dec(x_248); +lean_dec(x_247); +lean_dec_ref(x_246); +lean_dec(x_245); +lean_dec_ref(x_244); +lean_dec(x_243); lean_dec(x_242); lean_dec(x_241); lean_dec(x_240); lean_dec(x_239); lean_dec(x_238); lean_dec(x_237); -lean_dec(x_236); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -18645,68 +18116,68 @@ x_299 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_300 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_301 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_302 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_288); +lean_inc(x_285); x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_288); +lean_ctor_set(x_303, 0, x_285); lean_ctor_set(x_303, 1, x_301); lean_ctor_set(x_303, 2, x_302); lean_inc_ref(x_303); -lean_inc(x_288); -x_304 = l_Lean_Syntax_node1(x_288, x_300, x_303); +lean_inc(x_285); +x_304 = l_Lean_Syntax_node1(x_285, x_300, x_303); x_305 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_306 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_307 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_308 = l_Lean_addMacroScope(x_291, x_307, x_287); -lean_inc(x_286); -lean_inc(x_288); +x_308 = l_Lean_addMacroScope(x_291, x_307, x_288); +lean_inc(x_289); +lean_inc(x_285); x_309 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_309, 0, x_288); +lean_ctor_set(x_309, 0, x_285); lean_ctor_set(x_309, 1, x_306); lean_ctor_set(x_309, 2, x_308); -lean_ctor_set(x_309, 3, x_286); -lean_inc(x_288); -x_310 = l_Lean_Syntax_node2(x_288, x_305, x_309, x_303); -x_311 = l_Lean_Syntax_node2(x_288, x_299, x_304, x_310); +lean_ctor_set(x_309, 3, x_289); +lean_inc(x_285); +x_310 = l_Lean_Syntax_node2(x_285, x_305, x_309, x_303); +x_311 = l_Lean_Syntax_node2(x_285, x_299, x_304, x_310); if (lean_obj_tag(x_234) == 0) { lean_object* x_312; x_312 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_312); -x_236 = x_284; -x_237 = x_299; -x_238 = x_285; -x_239 = x_300; -x_240 = x_289; -x_241 = x_290; -x_242 = x_294; -x_243 = x_302; -x_244 = x_286; -x_245 = x_298; -x_246 = x_305; -x_247 = x_296; -x_248 = x_297; -x_249 = x_301; -x_250 = x_311; +x_236 = x_302; +x_237 = x_284; +x_238 = x_305; +x_239 = x_294; +x_240 = x_290; +x_241 = x_300; +x_242 = x_311; +x_243 = x_299; +x_244 = x_296; +x_245 = x_301; +x_246 = x_298; +x_247 = x_286; +x_248 = x_287; +x_249 = x_289; +x_250 = x_297; x_251 = lean_box(0); goto block_283; } else { -x_236 = x_284; -x_237 = x_299; -x_238 = x_285; -x_239 = x_300; -x_240 = x_289; -x_241 = x_290; -x_242 = x_294; -x_243 = x_302; -x_244 = x_286; -x_245 = x_298; -x_246 = x_305; -x_247 = x_296; -x_248 = x_297; -x_249 = x_301; -x_250 = x_311; +x_236 = x_302; +x_237 = x_284; +x_238 = x_305; +x_239 = x_294; +x_240 = x_290; +x_241 = x_300; +x_242 = x_311; +x_243 = x_299; +x_244 = x_296; +x_245 = x_301; +x_246 = x_298; +x_247 = x_286; +x_248 = x_287; +x_249 = x_289; +x_250 = x_297; x_251 = lean_box(0); goto block_283; } @@ -18846,13 +18317,13 @@ x_342 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lak x_343 = lean_ctor_get(x_342, 0); lean_inc(x_343); lean_dec_ref(x_342); -x_284 = x_331; -x_285 = x_324; -x_286 = x_328; -x_287 = x_341; -x_288 = x_339; -x_289 = x_337; -x_290 = x_329; +x_284 = x_329; +x_285 = x_339; +x_286 = x_324; +x_287 = x_337; +x_288 = x_341; +x_289 = x_328; +x_290 = x_331; x_291 = x_343; x_292 = lean_box(0); goto block_319; @@ -18865,13 +18336,13 @@ lean_inc(x_344); lean_dec_ref(x_340); x_345 = lean_ctor_get(x_234, 0); lean_inc(x_345); -x_284 = x_331; -x_285 = x_324; -x_286 = x_328; -x_287 = x_344; -x_288 = x_339; -x_289 = x_337; -x_290 = x_329; +x_284 = x_329; +x_285 = x_339; +x_286 = x_324; +x_287 = x_337; +x_288 = x_344; +x_289 = x_328; +x_290 = x_331; x_291 = x_345; x_292 = lean_box(0); goto block_319; @@ -19068,43 +18539,43 @@ return x_364; block_131: { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_14); -x_43 = l_Array_append___redArg(x_14, x_42); +lean_inc_ref(x_25); +x_43 = l_Array_append___redArg(x_25, x_42); lean_dec_ref(x_42); -lean_inc(x_20); -lean_inc(x_18); +lean_inc(x_21); +lean_inc(x_24); x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_18); -lean_ctor_set(x_44, 1, x_20); +lean_ctor_set(x_44, 0, x_24); +lean_ctor_set(x_44, 1, x_21); lean_ctor_set(x_44, 2, x_43); -lean_inc_n(x_29, 6); -lean_inc(x_13); -lean_inc(x_18); -x_45 = l_Lean_Syntax_node7(x_18, x_13, x_44, x_29, x_29, x_29, x_29, x_29, x_29); +lean_inc_n(x_33, 6); +lean_inc(x_28); +lean_inc(x_24); +x_45 = l_Lean_Syntax_node7(x_24, x_28, x_44, x_33, x_33, x_33, x_33, x_33, x_33); x_46 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; -lean_inc_ref(x_23); -lean_inc_ref(x_19); lean_inc_ref(x_35); -x_47 = l_Lean_Name_mkStr4(x_35, x_19, x_23, x_46); -lean_inc(x_18); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_47 = l_Lean_Name_mkStr4(x_20, x_40, x_35, x_46); +lean_inc(x_24); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_18); +lean_ctor_set(x_48, 0, x_24); lean_ctor_set(x_48, 1, x_46); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; -lean_inc_ref(x_23); -lean_inc_ref(x_19); lean_inc_ref(x_35); -x_50 = l_Lean_Name_mkStr4(x_35, x_19, x_23, x_49); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_50 = l_Lean_Name_mkStr4(x_20, x_40, x_35, x_49); x_51 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_52 = lean_box(2); -lean_inc(x_20); +lean_inc(x_21); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_20); +lean_ctor_set(x_53, 1, x_21); lean_ctor_set(x_53, 2, x_51); -x_54 = lean_mk_empty_array_with_capacity(x_36); -lean_inc(x_26); -x_55 = lean_array_push(x_54, x_26); +x_54 = lean_mk_empty_array_with_capacity(x_37); +lean_inc(x_22); +x_55 = lean_array_push(x_54, x_22); x_56 = lean_array_push(x_55, x_53); lean_inc(x_50); x_57 = lean_alloc_ctor(1, 3, 0); @@ -19112,194 +18583,194 @@ lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_50); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_23); -lean_inc_ref(x_19); lean_inc_ref(x_35); -x_59 = l_Lean_Name_mkStr4(x_35, x_19, x_23, x_58); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_59 = l_Lean_Name_mkStr4(x_20, x_40, x_35, x_58); x_60 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_17); -lean_inc_ref(x_19); -lean_inc_ref(x_35); -x_61 = l_Lean_Name_mkStr4(x_35, x_19, x_17, x_60); -lean_inc(x_15); +lean_inc_ref(x_34); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_61 = l_Lean_Name_mkStr4(x_20, x_40, x_34, x_60); +lean_inc(x_36); lean_inc(x_61); -lean_inc(x_18); -x_62 = l_Lean_Syntax_node2(x_18, x_61, x_15, x_12); -lean_inc(x_20); -lean_inc(x_18); -x_63 = l_Lean_Syntax_node1(x_18, x_20, x_62); -lean_inc(x_29); +lean_inc(x_24); +x_62 = l_Lean_Syntax_node2(x_24, x_61, x_36, x_31); +lean_inc(x_21); +lean_inc(x_24); +x_63 = l_Lean_Syntax_node1(x_24, x_21, x_62); +lean_inc(x_33); lean_inc(x_59); -lean_inc(x_18); -x_64 = l_Lean_Syntax_node2(x_18, x_59, x_29, x_63); +lean_inc(x_24); +x_64 = l_Lean_Syntax_node2(x_24, x_59, x_33, x_63); x_65 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__60; -lean_inc_ref(x_23); -lean_inc_ref(x_19); lean_inc_ref(x_35); -x_66 = l_Lean_Name_mkStr4(x_35, x_19, x_23, x_65); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_66 = l_Lean_Name_mkStr4(x_20, x_40, x_35, x_65); x_67 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__1; x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__2; -lean_inc_ref(x_21); -x_69 = l_Lean_Name_mkStr3(x_21, x_30, x_68); -lean_inc(x_39); +lean_inc_ref(x_23); +x_69 = l_Lean_Name_mkStr3(x_23, x_38, x_68); +lean_inc(x_12); lean_inc(x_69); -lean_inc(x_38); -x_70 = l_Lean_addMacroScope(x_38, x_69, x_39); -lean_inc(x_41); +lean_inc(x_30); +x_70 = l_Lean_addMacroScope(x_30, x_69, x_12); +lean_inc(x_14); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_41); -lean_inc(x_34); +lean_ctor_set(x_71, 1, x_14); +lean_inc(x_39); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_34); -lean_inc(x_18); +lean_ctor_set(x_72, 1, x_39); +lean_inc(x_24); x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_18); +lean_ctor_set(x_73, 0, x_24); lean_ctor_set(x_73, 1, x_67); lean_ctor_set(x_73, 2, x_70); lean_ctor_set(x_73, 3, x_72); -lean_inc(x_20); -lean_inc(x_18); -x_74 = l_Lean_Syntax_node4(x_18, x_20, x_32, x_24, x_28, x_33); -lean_inc(x_18); -x_75 = l_Lean_Syntax_node2(x_18, x_16, x_73, x_74); +lean_inc(x_21); +lean_inc(x_24); +x_74 = l_Lean_Syntax_node4(x_24, x_21, x_15, x_17, x_13, x_26); +lean_inc(x_24); +x_75 = l_Lean_Syntax_node2(x_24, x_16, x_73, x_74); x_76 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__62; x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__63; -lean_inc_ref(x_19); -lean_inc_ref(x_35); -x_78 = l_Lean_Name_mkStr4(x_35, x_19, x_76, x_77); -lean_inc_n(x_29, 2); -lean_inc(x_18); -x_79 = l_Lean_Syntax_node2(x_18, x_78, x_29, x_29); -lean_inc(x_29); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_78 = l_Lean_Name_mkStr4(x_20, x_40, x_76, x_77); +lean_inc_n(x_33, 2); +lean_inc(x_24); +x_79 = l_Lean_Syntax_node2(x_24, x_78, x_33, x_33); +lean_inc(x_33); lean_inc(x_79); -lean_inc(x_22); +lean_inc(x_41); lean_inc(x_66); -lean_inc(x_18); -x_80 = l_Lean_Syntax_node4(x_18, x_66, x_22, x_75, x_79, x_29); -lean_inc(x_18); -x_81 = l_Lean_Syntax_node4(x_18, x_47, x_48, x_57, x_64, x_80); -lean_inc(x_37); -lean_inc(x_18); -x_82 = l_Lean_Syntax_node2(x_18, x_37, x_45, x_81); +lean_inc(x_24); +x_80 = l_Lean_Syntax_node4(x_24, x_66, x_41, x_75, x_79, x_33); +lean_inc(x_24); +x_81 = l_Lean_Syntax_node4(x_24, x_47, x_48, x_57, x_64, x_80); +lean_inc(x_32); +lean_inc(x_24); +x_82 = l_Lean_Syntax_node2(x_24, x_32, x_45, x_81); x_83 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_17); -lean_inc_ref(x_19); -lean_inc_ref(x_35); -x_84 = l_Lean_Name_mkStr4(x_35, x_19, x_17, x_83); +lean_inc_ref(x_34); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_84 = l_Lean_Name_mkStr4(x_20, x_40, x_34, x_83); x_85 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_18); +lean_inc(x_24); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_18); +lean_ctor_set(x_86, 0, x_24); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Syntax_SepArray_ofElems(x_25, x_27); +x_87 = l_Lean_Syntax_SepArray_ofElems(x_19, x_27); lean_dec_ref(x_27); -x_88 = l_Array_append___redArg(x_14, x_87); +x_88 = l_Array_append___redArg(x_25, x_87); lean_dec_ref(x_87); -lean_inc(x_20); -lean_inc(x_18); +lean_inc(x_21); +lean_inc(x_24); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_18); -lean_ctor_set(x_89, 1, x_20); +lean_ctor_set(x_89, 0, x_24); +lean_ctor_set(x_89, 1, x_21); lean_ctor_set(x_89, 2, x_88); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_18); +lean_inc(x_24); x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_18); +lean_ctor_set(x_91, 0, x_24); lean_ctor_set(x_91, 1, x_90); -lean_inc(x_18); -x_92 = l_Lean_Syntax_node3(x_18, x_84, x_86, x_89, x_91); -lean_inc(x_20); -lean_inc(x_18); -x_93 = l_Lean_Syntax_node1(x_18, x_20, x_92); -lean_inc_n(x_29, 6); -lean_inc(x_18); -x_94 = l_Lean_Syntax_node7(x_18, x_13, x_29, x_93, x_29, x_29, x_29, x_29, x_29); +lean_inc(x_24); +x_92 = l_Lean_Syntax_node3(x_24, x_84, x_86, x_89, x_91); +lean_inc(x_21); +lean_inc(x_24); +x_93 = l_Lean_Syntax_node1(x_24, x_21, x_92); +lean_inc_n(x_33, 6); +lean_inc(x_24); +x_94 = l_Lean_Syntax_node7(x_24, x_28, x_33, x_93, x_33, x_33, x_33, x_33, x_33); x_95 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; -lean_inc_ref(x_19); -lean_inc_ref(x_35); -x_96 = l_Lean_Name_mkStr4(x_35, x_19, x_23, x_95); +lean_inc_ref(x_40); +lean_inc_ref(x_20); +x_96 = l_Lean_Name_mkStr4(x_20, x_40, x_35, x_95); x_97 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_18); +lean_inc(x_24); x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_18); +lean_ctor_set(x_98, 0, x_24); lean_ctor_set(x_98, 1, x_97); x_99 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_100 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; +lean_inc(x_12); +lean_inc(x_30); +x_101 = l_Lean_addMacroScope(x_30, x_100, x_12); lean_inc(x_39); -lean_inc(x_38); -x_101 = l_Lean_addMacroScope(x_38, x_100, x_39); -lean_inc(x_34); -lean_inc(x_18); +lean_inc(x_24); x_102 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_102, 0, x_18); +lean_ctor_set(x_102, 0, x_24); lean_ctor_set(x_102, 1, x_99); lean_ctor_set(x_102, 2, x_101); -lean_ctor_set(x_102, 3, x_34); -lean_inc(x_29); -lean_inc(x_18); -x_103 = l_Lean_Syntax_node2(x_18, x_50, x_102, x_29); +lean_ctor_set(x_102, 3, x_39); +lean_inc(x_33); +lean_inc(x_24); +x_103 = l_Lean_Syntax_node2(x_24, x_50, x_102, x_33); x_104 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__5; x_105 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_39); -lean_inc(x_38); -x_107 = l_Lean_addMacroScope(x_38, x_106, x_39); -x_108 = l_Lean_Name_mkStr2(x_21, x_104); +lean_inc(x_12); +lean_inc(x_30); +x_107 = l_Lean_addMacroScope(x_30, x_106, x_12); +x_108 = l_Lean_Name_mkStr2(x_23, x_104); lean_inc(x_108); x_109 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_41); +lean_ctor_set(x_109, 1, x_14); x_110 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_110, 0, x_108); -lean_inc(x_34); +lean_inc(x_39); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_34); +lean_ctor_set(x_111, 1, x_39); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_109); lean_ctor_set(x_112, 1, x_111); -lean_inc(x_18); +lean_inc(x_24); x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_18); +lean_ctor_set(x_113, 0, x_24); lean_ctor_set(x_113, 1, x_105); lean_ctor_set(x_113, 2, x_107); lean_ctor_set(x_113, 3, x_112); -lean_inc(x_18); -x_114 = l_Lean_Syntax_node2(x_18, x_61, x_15, x_113); -lean_inc(x_20); -lean_inc(x_18); -x_115 = l_Lean_Syntax_node1(x_18, x_20, x_114); -lean_inc(x_29); -lean_inc(x_18); -x_116 = l_Lean_Syntax_node2(x_18, x_59, x_29, x_115); +lean_inc(x_24); +x_114 = l_Lean_Syntax_node2(x_24, x_61, x_36, x_113); +lean_inc(x_21); +lean_inc(x_24); +x_115 = l_Lean_Syntax_node1(x_24, x_21, x_114); +lean_inc(x_33); +lean_inc(x_24); +x_116 = l_Lean_Syntax_node2(x_24, x_59, x_33, x_115); x_117 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_118 = l_Lean_Name_mkStr4(x_35, x_19, x_17, x_117); +x_118 = l_Lean_Name_mkStr4(x_20, x_40, x_34, x_117); x_119 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_18); +lean_inc(x_24); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_18); +lean_ctor_set(x_120, 0, x_24); lean_ctor_set(x_120, 1, x_119); x_121 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_123 = l_Lean_addMacroScope(x_38, x_122, x_39); -lean_inc(x_18); +x_123 = l_Lean_addMacroScope(x_30, x_122, x_12); +lean_inc(x_24); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_18); +lean_ctor_set(x_124, 0, x_24); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_34); -lean_inc(x_18); -x_125 = l_Lean_Syntax_node3(x_18, x_118, x_26, x_120, x_124); -lean_inc(x_29); -lean_inc(x_18); -x_126 = l_Lean_Syntax_node4(x_18, x_66, x_22, x_125, x_79, x_29); -lean_inc(x_18); -x_127 = l_Lean_Syntax_node5(x_18, x_96, x_98, x_103, x_116, x_126, x_29); -lean_inc(x_18); -x_128 = l_Lean_Syntax_node2(x_18, x_37, x_94, x_127); -x_129 = l_Lean_Syntax_node3(x_18, x_20, x_40, x_82, x_128); +lean_ctor_set(x_124, 3, x_39); +lean_inc(x_24); +x_125 = l_Lean_Syntax_node3(x_24, x_118, x_22, x_120, x_124); +lean_inc(x_33); +lean_inc(x_24); +x_126 = l_Lean_Syntax_node4(x_24, x_66, x_41, x_125, x_79, x_33); +lean_inc(x_24); +x_127 = l_Lean_Syntax_node5(x_24, x_96, x_98, x_103, x_116, x_126, x_33); +lean_inc(x_24); +x_128 = l_Lean_Syntax_node2(x_24, x_32, x_94, x_127); +x_129 = l_Lean_Syntax_node3(x_24, x_21, x_18, x_82, x_128); x_130 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_130, 0, x_129); return x_130; @@ -19309,67 +18780,67 @@ return x_130; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; x_149 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__0; x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; -lean_inc_ref(x_138); -lean_inc(x_144); -lean_inc(x_142); +lean_inc_ref(x_133); +lean_inc(x_141); +lean_inc(x_132); x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_142); -lean_ctor_set(x_151, 1, x_144); -lean_ctor_set(x_151, 2, x_138); +lean_ctor_set(x_151, 0, x_132); +lean_ctor_set(x_151, 1, x_141); +lean_ctor_set(x_151, 2, x_133); x_152 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; -lean_inc(x_142); +lean_inc(x_132); x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_142); +lean_ctor_set(x_153, 0, x_132); lean_ctor_set(x_153, 1, x_152); x_154 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_142); +lean_inc(x_132); x_155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_155, 0, x_142); +lean_ctor_set(x_155, 0, x_132); lean_ctor_set(x_155, 1, x_154); x_156 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_157 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_146); +lean_inc(x_134); lean_inc(x_147); -x_158 = l_Lean_addMacroScope(x_147, x_157, x_146); +x_158 = l_Lean_addMacroScope(x_147, x_157, x_134); x_159 = lean_box(0); x_160 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__3; -lean_inc(x_139); +lean_inc(x_145); x_161 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_139); -lean_inc(x_142); +lean_ctor_set(x_161, 1, x_145); +lean_inc(x_132); x_162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_162, 0, x_142); +lean_ctor_set(x_162, 0, x_132); lean_ctor_set(x_162, 1, x_156); lean_ctor_set(x_162, 2, x_158); lean_ctor_set(x_162, 3, x_161); x_163 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_141); -lean_inc_ref(x_145); +lean_inc_ref(x_142); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_164 = l_Lean_Name_mkStr4(x_140, x_145, x_141, x_163); +x_164 = l_Lean_Name_mkStr4(x_140, x_146, x_142, x_163); x_165 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_141); -lean_inc_ref(x_145); +lean_inc_ref(x_142); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_166 = l_Lean_Name_mkStr4(x_140, x_145, x_141, x_165); +x_166 = l_Lean_Name_mkStr4(x_140, x_146, x_142, x_165); x_167 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_142); +lean_inc(x_132); x_168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_168, 0, x_142); +lean_ctor_set(x_168, 0, x_132); lean_ctor_set(x_168, 1, x_167); x_169 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_170 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_171 = lean_box(0); -lean_inc(x_146); +lean_inc(x_134); lean_inc(x_147); -x_172 = l_Lean_addMacroScope(x_147, x_171, x_146); +x_172 = l_Lean_addMacroScope(x_147, x_171, x_134); x_173 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__1; x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__59; -lean_inc_ref(x_145); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_176 = l_Lean_Name_mkStr3(x_140, x_145, x_175); +x_176 = l_Lean_Name_mkStr3(x_140, x_146, x_175); x_177 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_177, 0, x_176); x_178 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; @@ -19381,19 +18852,19 @@ lean_inc_ref(x_140); x_181 = l_Lean_Name_mkStr2(x_140, x_178); x_182 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_182, 0, x_181); -lean_inc_ref(x_145); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_183 = l_Lean_Name_mkStr2(x_140, x_145); +x_183 = l_Lean_Name_mkStr2(x_140, x_146); x_184 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_184, 0, x_183); lean_inc_ref(x_140); x_185 = l_Lean_Name_mkStr1(x_140); x_186 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_186, 0, x_185); -lean_inc(x_139); +lean_inc(x_145); x_187 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_139); +lean_ctor_set(x_187, 1, x_145); x_188 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_188, 0, x_184); lean_ctor_set(x_188, 1, x_187); @@ -19409,129 +18880,129 @@ lean_ctor_set(x_191, 1, x_190); x_192 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_192, 0, x_174); lean_ctor_set(x_192, 1, x_191); -lean_inc(x_142); +lean_inc(x_132); x_193 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_193, 0, x_142); +lean_ctor_set(x_193, 0, x_132); lean_ctor_set(x_193, 1, x_170); lean_ctor_set(x_193, 2, x_172); lean_ctor_set(x_193, 3, x_192); -lean_inc(x_142); -x_194 = l_Lean_Syntax_node1(x_142, x_169, x_193); -lean_inc(x_142); -x_195 = l_Lean_Syntax_node2(x_142, x_166, x_168, x_194); +lean_inc(x_132); +x_194 = l_Lean_Syntax_node1(x_132, x_169, x_193); +lean_inc(x_132); +x_195 = l_Lean_Syntax_node2(x_132, x_166, x_168, x_194); x_196 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__37; x_197 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__38; -lean_inc(x_142); +lean_inc(x_132); x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_142); +lean_ctor_set(x_198, 0, x_132); lean_ctor_set(x_198, 1, x_197); -lean_inc(x_142); -x_199 = l_Lean_Syntax_node1(x_142, x_196, x_198); +lean_inc(x_132); +x_199 = l_Lean_Syntax_node1(x_132, x_196, x_198); x_200 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -lean_inc(x_142); +lean_inc(x_132); x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_142); +lean_ctor_set(x_201, 0, x_132); lean_ctor_set(x_201, 1, x_200); +lean_inc(x_138); +lean_inc(x_141); lean_inc(x_132); -lean_inc(x_144); -lean_inc(x_142); -x_202 = l_Lean_Syntax_node1(x_142, x_144, x_132); +x_202 = l_Lean_Syntax_node1(x_132, x_141, x_138); lean_inc(x_199); -lean_inc(x_144); -lean_inc(x_142); -x_203 = l_Lean_Syntax_node3(x_142, x_144, x_199, x_201, x_202); +lean_inc(x_141); +lean_inc(x_132); +x_203 = l_Lean_Syntax_node3(x_132, x_141, x_199, x_201, x_202); x_204 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_142); +lean_inc(x_132); x_205 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_205, 0, x_142); +lean_ctor_set(x_205, 0, x_132); lean_ctor_set(x_205, 1, x_204); -lean_inc(x_142); -x_206 = l_Lean_Syntax_node3(x_142, x_164, x_195, x_203, x_205); +lean_inc(x_132); +x_206 = l_Lean_Syntax_node3(x_132, x_164, x_195, x_203, x_205); x_207 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_142); +lean_inc(x_132); x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_142); +lean_ctor_set(x_208, 0, x_132); lean_ctor_set(x_208, 1, x_207); x_209 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_141); -lean_inc_ref(x_145); +lean_inc_ref(x_142); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_210 = l_Lean_Name_mkStr4(x_140, x_145, x_141, x_209); +x_210 = l_Lean_Name_mkStr4(x_140, x_146, x_142, x_209); x_211 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_212 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_146); +lean_inc(x_134); lean_inc(x_147); -x_213 = l_Lean_addMacroScope(x_147, x_212, x_146); +x_213 = l_Lean_addMacroScope(x_147, x_212, x_134); x_214 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__8; x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_139); +lean_inc(x_145); x_216 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_139); +lean_ctor_set(x_216, 1, x_145); x_217 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_217, 0, x_214); lean_ctor_set(x_217, 1, x_216); -lean_inc(x_142); +lean_inc(x_132); x_218 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_218, 0, x_142); +lean_ctor_set(x_218, 0, x_132); lean_ctor_set(x_218, 1, x_211); lean_ctor_set(x_218, 2, x_213); lean_ctor_set(x_218, 3, x_217); -lean_inc(x_135); -lean_inc(x_144); -lean_inc(x_142); -x_219 = l_Lean_Syntax_node1(x_142, x_144, x_135); +lean_inc(x_136); +lean_inc(x_141); +lean_inc(x_132); +x_219 = l_Lean_Syntax_node1(x_132, x_141, x_136); lean_inc(x_210); -lean_inc(x_142); -x_220 = l_Lean_Syntax_node2(x_142, x_210, x_218, x_219); +lean_inc(x_132); +x_220 = l_Lean_Syntax_node2(x_132, x_210, x_218, x_219); lean_inc_ref(x_208); lean_inc_ref(x_155); -lean_inc(x_133); +lean_inc(x_143); lean_inc_ref(x_151); -lean_inc(x_142); -x_221 = l_Lean_Syntax_node8(x_142, x_150, x_151, x_153, x_133, x_155, x_162, x_206, x_208, x_220); +lean_inc(x_132); +x_221 = l_Lean_Syntax_node8(x_132, x_150, x_151, x_153, x_143, x_155, x_162, x_206, x_208, x_220); x_222 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_145); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_223 = l_Lean_Name_mkStr4(x_140, x_145, x_175, x_222); +x_223 = l_Lean_Name_mkStr4(x_140, x_146, x_175, x_222); x_224 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_145); +lean_inc_ref(x_146); lean_inc_ref(x_140); -x_225 = l_Lean_Name_mkStr4(x_140, x_145, x_175, x_224); +x_225 = l_Lean_Name_mkStr4(x_140, x_146, x_175, x_224); if (lean_obj_tag(x_5) == 0) { lean_object* x_226; x_226 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_12 = x_136; -x_13 = x_225; -x_14 = x_138; -x_15 = x_155; +x_12 = x_134; +x_13 = x_136; +x_14 = x_159; +x_15 = x_199; x_16 = x_210; -x_17 = x_141; -x_18 = x_142; -x_19 = x_145; -x_20 = x_144; -x_21 = x_149; -x_22 = x_208; -x_23 = x_175; +x_17 = x_138; +x_18 = x_221; +x_19 = x_200; +x_20 = x_140; +x_21 = x_141; +x_22 = x_143; +x_23 = x_149; x_24 = x_132; -x_25 = x_200; -x_26 = x_133; -x_27 = x_134; -x_28 = x_135; -x_29 = x_151; -x_30 = x_173; -x_31 = lean_box(0); -x_32 = x_199; -x_33 = x_137; -x_34 = x_139; -x_35 = x_140; -x_36 = x_143; -x_37 = x_223; -x_38 = x_147; -x_39 = x_146; -x_40 = x_221; -x_41 = x_159; +x_25 = x_133; +x_26 = x_135; +x_27 = x_137; +x_28 = x_225; +x_29 = lean_box(0); +x_30 = x_147; +x_31 = x_139; +x_32 = x_223; +x_33 = x_151; +x_34 = x_142; +x_35 = x_175; +x_36 = x_155; +x_37 = x_144; +x_38 = x_173; +x_39 = x_145; +x_40 = x_146; +x_41 = x_208; x_42 = x_226; goto block_131; } @@ -19542,36 +19013,36 @@ x_227 = lean_ctor_get(x_5, 0); lean_inc(x_227); lean_dec_ref(x_5); x_228 = l_Array_mkArray1___redArg(x_227); -x_12 = x_136; -x_13 = x_225; -x_14 = x_138; -x_15 = x_155; +x_12 = x_134; +x_13 = x_136; +x_14 = x_159; +x_15 = x_199; x_16 = x_210; -x_17 = x_141; -x_18 = x_142; -x_19 = x_145; -x_20 = x_144; -x_21 = x_149; -x_22 = x_208; -x_23 = x_175; +x_17 = x_138; +x_18 = x_221; +x_19 = x_200; +x_20 = x_140; +x_21 = x_141; +x_22 = x_143; +x_23 = x_149; x_24 = x_132; -x_25 = x_200; -x_26 = x_133; -x_27 = x_134; -x_28 = x_135; -x_29 = x_151; -x_30 = x_173; -x_31 = lean_box(0); -x_32 = x_199; -x_33 = x_137; -x_34 = x_139; -x_35 = x_140; -x_36 = x_143; -x_37 = x_223; -x_38 = x_147; -x_39 = x_146; -x_40 = x_221; -x_41 = x_159; +x_25 = x_133; +x_26 = x_135; +x_27 = x_137; +x_28 = x_225; +x_29 = lean_box(0); +x_30 = x_147; +x_31 = x_139; +x_32 = x_223; +x_33 = x_151; +x_34 = x_142; +x_35 = x_175; +x_36 = x_155; +x_37 = x_144; +x_38 = x_173; +x_39 = x_145; +x_40 = x_146; +x_41 = x_208; x_42 = x_228; goto block_131; } @@ -19716,7 +19187,7 @@ lean_dec(x_14); lean_ctor_set(x_2, 7, x_29); lean_inc(x_3); lean_inc_ref(x_2); -x_30 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0(x_27, x_28, x_28, x_26, x_21, x_20, x_19, x_18, x_2, x_3); +x_30 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0(x_27, x_28, x_28, x_26, x_21, x_19, x_20, x_18, x_2, x_3); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -19795,7 +19266,7 @@ lean_ctor_set(x_51, 9, x_45); lean_ctor_set_uint8(x_51, sizeof(void*)*10, x_46); lean_inc(x_3); lean_inc_ref(x_51); -x_52 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0(x_48, x_49, x_49, x_47, x_21, x_20, x_19, x_18, x_51, x_3); +x_52 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0(x_48, x_49, x_49, x_47, x_21, x_19, x_20, x_18, x_51, x_3); if (lean_obj_tag(x_52) == 0) { lean_object* x_53; lean_object* x_54; lean_object* x_55; @@ -19870,8 +19341,8 @@ if (lean_obj_tag(x_65) == 0) { lean_object* x_66; x_66 = lean_box(0); -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_66; goto block_62; } @@ -19881,8 +19352,8 @@ uint8_t x_67; x_67 = !lean_is_exclusive(x_65); if (x_67 == 0) { -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_65; goto block_62; } @@ -19894,8 +19365,8 @@ lean_inc(x_68); lean_dec(x_65); x_69 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_69, 0, x_68); -x_19 = x_63; -x_20 = x_64; +x_19 = x_64; +x_20 = x_63; x_21 = x_69; goto block_62; } @@ -19941,15 +19412,6 @@ goto block_70; } } } -LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___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_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_spec__0_spec__0___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_7); -return x_12; -} -} LEAN_EXPORT lean_object* l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabInputfileCommand_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) { _start: { @@ -20097,16 +19559,16 @@ lean_ctor_set(x_30, 0, x_19); lean_ctor_set(x_30, 1, x_29); x_31 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__4; x_32 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; -lean_inc(x_20); +lean_inc(x_18); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 0, x_18); lean_ctor_set(x_33, 1, x_24); lean_ctor_set(x_33, 2, x_32); x_34 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; x_35 = lean_array_push(x_34, x_3); x_36 = lean_array_push(x_35, x_33); x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_20); +lean_ctor_set(x_37, 0, x_18); lean_ctor_set(x_37, 1, x_31); lean_ctor_set(x_37, 2, x_36); x_38 = l_Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6___lam__0___closed__5; @@ -20124,7 +19586,7 @@ lean_inc_ref(x_26); lean_inc(x_19); x_44 = l_Lean_Syntax_node2(x_19, x_38, x_26, x_43); lean_inc(x_19); -x_45 = l_Lean_Syntax_node5(x_19, x_28, x_30, x_37, x_44, x_18, x_26); +x_45 = l_Lean_Syntax_node5(x_19, x_28, x_30, x_37, x_44, x_20, x_26); x_46 = l_Lean_Syntax_node2(x_19, x_22, x_27, x_45); lean_inc(x_46); x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___boxed), 4, 1); @@ -20166,17 +19628,17 @@ if (lean_obj_tag(x_55) == 0) lean_object* x_65; x_65 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_65); -x_18 = x_62; +x_18 = x_61; x_19 = x_64; -x_20 = x_61; +x_20 = x_62; x_21 = lean_box(0); goto block_49; } else { -x_18 = x_62; +x_18 = x_61; x_19 = x_64; -x_20 = x_61; +x_20 = x_62; x_21 = lean_box(0); goto block_49; } @@ -20683,22 +20145,22 @@ lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); lean_dec_ref(x_258); -lean_inc_ref(x_248); -lean_inc(x_242); -lean_inc(x_241); +lean_inc_ref(x_241); +lean_inc(x_247); +lean_inc(x_244); x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_241); -lean_ctor_set(x_260, 1, x_242); -lean_ctor_set(x_260, 2, x_248); +lean_ctor_set(x_260, 0, x_244); +lean_ctor_set(x_260, 1, x_247); +lean_ctor_set(x_260, 2, x_241); lean_inc_ref(x_260); -lean_inc(x_241); -x_261 = l_Lean_Syntax_node1(x_241, x_247, x_260); -lean_inc(x_241); -x_262 = l_Lean_Syntax_node2(x_241, x_244, x_238, x_260); -x_263 = l_Lean_Syntax_node2(x_241, x_236, x_261, x_262); +lean_inc(x_244); +x_261 = l_Lean_Syntax_node1(x_244, x_237, x_260); +lean_inc(x_244); +x_262 = l_Lean_Syntax_node2(x_244, x_250, x_238, x_260); +x_263 = l_Lean_Syntax_node2(x_244, x_245, x_261, x_262); x_264 = lean_unsigned_to_nat(2u); x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -x_266 = lean_array_push(x_265, x_250); +x_266 = lean_array_push(x_265, x_240); x_267 = lean_array_push(x_266, x_263); x_268 = l_Lake_DSL_expandAttrs(x_6); x_269 = l_Array_append___redArg(x_267, x_268); @@ -20712,21 +20174,21 @@ lean_dec(x_10); x_272 = lean_ctor_get(x_271, 0); lean_inc(x_272); lean_dec_ref(x_271); -x_132 = x_269; -x_133 = x_237; -x_134 = x_264; -x_135 = x_239; -x_136 = x_240; -x_137 = x_253; -x_138 = x_242; -x_139 = x_243; -x_140 = x_257; -x_141 = x_245; -x_142 = x_246; -x_143 = x_259; -x_144 = x_270; -x_145 = x_248; -x_146 = x_249; +x_132 = x_236; +x_133 = x_239; +x_134 = x_257; +x_135 = x_270; +x_136 = x_241; +x_137 = x_242; +x_138 = x_243; +x_139 = x_269; +x_140 = x_264; +x_141 = x_247; +x_142 = x_248; +x_143 = x_246; +x_144 = x_249; +x_145 = x_253; +x_146 = x_259; x_147 = x_272; x_148 = lean_box(0); goto block_229; @@ -20738,21 +20200,21 @@ lean_dec(x_10); x_273 = lean_ctor_get(x_234, 0); lean_inc(x_273); lean_dec_ref(x_234); -x_132 = x_269; -x_133 = x_237; -x_134 = x_264; -x_135 = x_239; -x_136 = x_240; -x_137 = x_253; -x_138 = x_242; -x_139 = x_243; -x_140 = x_257; -x_141 = x_245; -x_142 = x_246; -x_143 = x_259; -x_144 = x_270; -x_145 = x_248; -x_146 = x_249; +x_132 = x_236; +x_133 = x_239; +x_134 = x_257; +x_135 = x_270; +x_136 = x_241; +x_137 = x_242; +x_138 = x_243; +x_139 = x_269; +x_140 = x_264; +x_141 = x_247; +x_142 = x_248; +x_143 = x_246; +x_144 = x_249; +x_145 = x_253; +x_146 = x_259; x_147 = x_273; x_148 = lean_box(0); goto block_229; @@ -20765,7 +20227,7 @@ lean_dec(x_257); lean_dec(x_255); lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec_ref(x_248); lean_dec(x_247); lean_dec(x_246); @@ -20773,12 +20235,12 @@ lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); lean_dec(x_242); -lean_dec(x_241); +lean_dec_ref(x_241); lean_dec(x_240); lean_dec_ref(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec(x_236); +lean_dec(x_237); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec(x_6); @@ -20807,7 +20269,7 @@ uint8_t x_277; lean_dec(x_255); lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec_ref(x_248); lean_dec(x_247); lean_dec(x_246); @@ -20815,12 +20277,12 @@ lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); lean_dec(x_242); -lean_dec(x_241); +lean_dec_ref(x_241); lean_dec(x_240); lean_dec_ref(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec(x_236); +lean_dec(x_237); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -20849,7 +20311,7 @@ else uint8_t x_280; lean_dec(x_253); lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec_ref(x_248); lean_dec(x_247); lean_dec(x_246); @@ -20857,12 +20319,12 @@ lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); lean_dec(x_242); -lean_dec(x_241); +lean_dec_ref(x_241); lean_dec(x_240); lean_dec_ref(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec(x_236); +lean_dec(x_237); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -20889,7 +20351,7 @@ return x_282; else { lean_dec(x_250); -lean_dec_ref(x_249); +lean_dec(x_249); lean_dec_ref(x_248); lean_dec(x_247); lean_dec(x_246); @@ -20897,12 +20359,12 @@ lean_dec(x_245); lean_dec(x_244); lean_dec(x_243); lean_dec(x_242); -lean_dec(x_241); +lean_dec_ref(x_241); lean_dec(x_240); lean_dec_ref(x_239); lean_dec(x_238); -lean_dec_ref(x_237); -lean_dec(x_236); +lean_dec(x_237); +lean_dec_ref(x_236); lean_dec(x_234); lean_dec(x_10); lean_dec_ref(x_9); @@ -20934,68 +20396,68 @@ x_299 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__14; x_300 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__15; x_301 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_302 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; -lean_inc(x_289); +lean_inc(x_287); x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_289); +lean_ctor_set(x_303, 0, x_287); lean_ctor_set(x_303, 1, x_301); lean_ctor_set(x_303, 2, x_302); lean_inc_ref(x_303); -lean_inc(x_289); -x_304 = l_Lean_Syntax_node1(x_289, x_300, x_303); +lean_inc(x_287); +x_304 = l_Lean_Syntax_node1(x_287, x_300, x_303); x_305 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__16; x_306 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_307 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; -x_308 = l_Lean_addMacroScope(x_291, x_307, x_287); -lean_inc(x_284); -lean_inc(x_289); +x_308 = l_Lean_addMacroScope(x_291, x_307, x_289); +lean_inc(x_285); +lean_inc(x_287); x_309 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_309, 0, x_289); +lean_ctor_set(x_309, 0, x_287); lean_ctor_set(x_309, 1, x_306); lean_ctor_set(x_309, 2, x_308); -lean_ctor_set(x_309, 3, x_284); -lean_inc(x_289); -x_310 = l_Lean_Syntax_node2(x_289, x_305, x_309, x_303); -x_311 = l_Lean_Syntax_node2(x_289, x_299, x_304, x_310); +lean_ctor_set(x_309, 3, x_285); +lean_inc(x_287); +x_310 = l_Lean_Syntax_node2(x_287, x_305, x_309, x_303); +x_311 = l_Lean_Syntax_node2(x_287, x_299, x_304, x_310); if (lean_obj_tag(x_234) == 0) { lean_object* x_312; x_312 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___at___00__private_Lake_DSL_Targets_0__Lake_DSL_elabLeanLibCommand_spec__6_spec__6_spec__6___redArg(x_10); lean_dec_ref(x_312); -x_236 = x_299; -x_237 = x_297; +x_236 = x_297; +x_237 = x_300; x_238 = x_286; x_239 = x_298; -x_240 = x_288; -x_241 = x_294; -x_242 = x_301; -x_243 = x_290; -x_244 = x_305; -x_245 = x_284; -x_246 = x_285; -x_247 = x_300; -x_248 = x_302; -x_249 = x_296; -x_250 = x_311; +x_240 = x_311; +x_241 = x_302; +x_242 = x_284; +x_243 = x_285; +x_244 = x_294; +x_245 = x_299; +x_246 = x_288; +x_247 = x_301; +x_248 = x_296; +x_249 = x_290; +x_250 = x_305; x_251 = lean_box(0); goto block_283; } else { -x_236 = x_299; -x_237 = x_297; +x_236 = x_297; +x_237 = x_300; x_238 = x_286; x_239 = x_298; -x_240 = x_288; -x_241 = x_294; -x_242 = x_301; -x_243 = x_290; -x_244 = x_305; -x_245 = x_284; -x_246 = x_285; -x_247 = x_300; -x_248 = x_302; -x_249 = x_296; -x_250 = x_311; +x_240 = x_311; +x_241 = x_302; +x_242 = x_284; +x_243 = x_285; +x_244 = x_294; +x_245 = x_299; +x_246 = x_288; +x_247 = x_301; +x_248 = x_296; +x_249 = x_290; +x_250 = x_305; x_251 = lean_box(0); goto block_283; } @@ -21135,13 +20597,13 @@ x_342 = l_Lean_getMainModule___at___00Lake_DSL_elabConfig___at___00__private_Lak x_343 = lean_ctor_get(x_342, 0); lean_inc(x_343); lean_dec_ref(x_342); -x_284 = x_328; -x_285 = x_324; +x_284 = x_329; +x_285 = x_328; x_286 = x_337; -x_287 = x_341; +x_287 = x_339; x_288 = x_331; -x_289 = x_339; -x_290 = x_329; +x_289 = x_341; +x_290 = x_324; x_291 = x_343; x_292 = lean_box(0); goto block_319; @@ -21154,13 +20616,13 @@ lean_inc(x_344); lean_dec_ref(x_340); x_345 = lean_ctor_get(x_234, 0); lean_inc(x_345); -x_284 = x_328; -x_285 = x_324; +x_284 = x_329; +x_285 = x_328; x_286 = x_337; -x_287 = x_344; +x_287 = x_339; x_288 = x_331; -x_289 = x_339; -x_290 = x_329; +x_289 = x_344; +x_290 = x_324; x_291 = x_345; x_292 = lean_box(0); goto block_319; @@ -21357,43 +20819,43 @@ return x_364; block_131: { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_24); -x_43 = l_Array_append___redArg(x_24, x_42); +lean_inc_ref(x_31); +x_43 = l_Array_append___redArg(x_31, x_42); lean_dec_ref(x_42); -lean_inc(x_35); -lean_inc(x_38); +lean_inc(x_22); +lean_inc(x_15); x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_38); -lean_ctor_set(x_44, 1, x_35); +lean_ctor_set(x_44, 0, x_15); +lean_ctor_set(x_44, 1, x_22); lean_ctor_set(x_44, 2, x_43); -lean_inc_n(x_13, 6); -lean_inc(x_39); -lean_inc(x_38); -x_45 = l_Lean_Syntax_node7(x_38, x_39, x_44, x_13, x_13, x_13, x_13, x_13, x_13); +lean_inc_n(x_33, 6); +lean_inc(x_25); +lean_inc(x_15); +x_45 = l_Lean_Syntax_node7(x_15, x_25, x_44, x_33, x_33, x_33, x_33, x_33, x_33); x_46 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__9; +lean_inc_ref(x_28); lean_inc_ref(x_12); -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_47 = l_Lean_Name_mkStr4(x_26, x_31, x_12, x_46); -lean_inc(x_38); +lean_inc_ref(x_38); +x_47 = l_Lean_Name_mkStr4(x_38, x_12, x_28, x_46); +lean_inc(x_15); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_38); +lean_ctor_set(x_48, 0, x_15); lean_ctor_set(x_48, 1, x_46); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; +lean_inc_ref(x_28); lean_inc_ref(x_12); -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_50 = l_Lean_Name_mkStr4(x_26, x_31, x_12, x_49); +lean_inc_ref(x_38); +x_50 = l_Lean_Name_mkStr4(x_38, x_12, x_28, x_49); x_51 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_52 = lean_box(2); -lean_inc(x_35); +lean_inc(x_22); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_35); +lean_ctor_set(x_53, 1, x_22); lean_ctor_set(x_53, 2, x_51); -x_54 = lean_mk_empty_array_with_capacity(x_33); -lean_inc(x_21); -x_55 = lean_array_push(x_54, x_21); +x_54 = lean_mk_empty_array_with_capacity(x_20); +lean_inc(x_23); +x_55 = lean_array_push(x_54, x_23); x_56 = lean_array_push(x_55, x_53); lean_inc(x_50); x_57 = lean_alloc_ctor(1, 3, 0); @@ -21401,194 +20863,194 @@ lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_50); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; +lean_inc_ref(x_28); lean_inc_ref(x_12); -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_59 = l_Lean_Name_mkStr4(x_26, x_31, x_12, x_58); +lean_inc_ref(x_38); +x_59 = l_Lean_Name_mkStr4(x_38, x_12, x_28, x_58); x_60 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_17); -lean_inc_ref(x_31); lean_inc_ref(x_26); -x_61 = l_Lean_Name_mkStr4(x_26, x_31, x_17, x_60); -lean_inc(x_34); -lean_inc(x_61); -lean_inc(x_38); -x_62 = l_Lean_Syntax_node2(x_38, x_61, x_34, x_19); -lean_inc(x_35); -lean_inc(x_38); -x_63 = l_Lean_Syntax_node1(x_38, x_35, x_62); +lean_inc_ref(x_12); +lean_inc_ref(x_38); +x_61 = l_Lean_Name_mkStr4(x_38, x_12, x_26, x_60); lean_inc(x_13); +lean_inc(x_61); +lean_inc(x_15); +x_62 = l_Lean_Syntax_node2(x_15, x_61, x_13, x_40); +lean_inc(x_22); +lean_inc(x_15); +x_63 = l_Lean_Syntax_node1(x_15, x_22, x_62); +lean_inc(x_33); lean_inc(x_59); -lean_inc(x_38); -x_64 = l_Lean_Syntax_node2(x_38, x_59, x_13, x_63); +lean_inc(x_15); +x_64 = l_Lean_Syntax_node2(x_15, x_59, x_33, x_63); x_65 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__60; +lean_inc_ref(x_28); lean_inc_ref(x_12); -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_66 = l_Lean_Name_mkStr4(x_26, x_31, x_12, x_65); +lean_inc_ref(x_38); +x_66 = l_Lean_Name_mkStr4(x_38, x_12, x_28, x_65); x_67 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__1; x_68 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__2; -lean_inc_ref(x_15); -x_69 = l_Lean_Name_mkStr3(x_15, x_16, x_68); -lean_inc(x_22); +lean_inc_ref(x_29); +x_69 = l_Lean_Name_mkStr3(x_29, x_41, x_68); +lean_inc(x_24); lean_inc(x_69); +lean_inc(x_34); +x_70 = l_Lean_addMacroScope(x_34, x_69, x_24); lean_inc(x_37); -x_70 = l_Lean_addMacroScope(x_37, x_69, x_22); -lean_inc(x_27); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_27); -lean_inc(x_20); +lean_ctor_set(x_71, 1, x_37); +lean_inc(x_17); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_20); -lean_inc(x_38); +lean_ctor_set(x_72, 1, x_17); +lean_inc(x_15); x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_38); +lean_ctor_set(x_73, 0, x_15); lean_ctor_set(x_73, 1, x_67); lean_ctor_set(x_73, 2, x_70); lean_ctor_set(x_73, 3, x_72); -lean_inc(x_35); -lean_inc(x_38); -x_74 = l_Lean_Syntax_node4(x_38, x_35, x_25, x_18, x_23, x_36); -lean_inc(x_38); -x_75 = l_Lean_Syntax_node2(x_38, x_32, x_73, x_74); +lean_inc(x_22); +lean_inc(x_15); +x_74 = l_Lean_Syntax_node4(x_15, x_22, x_14, x_39, x_30, x_32); +lean_inc(x_15); +x_75 = l_Lean_Syntax_node2(x_15, x_18, x_73, x_74); x_76 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__62; x_77 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__63; -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_78 = l_Lean_Name_mkStr4(x_26, x_31, x_76, x_77); -lean_inc_n(x_13, 2); -lean_inc(x_38); -x_79 = l_Lean_Syntax_node2(x_38, x_78, x_13, x_13); -lean_inc(x_13); +lean_inc_ref(x_12); +lean_inc_ref(x_38); +x_78 = l_Lean_Name_mkStr4(x_38, x_12, x_76, x_77); +lean_inc_n(x_33, 2); +lean_inc(x_15); +x_79 = l_Lean_Syntax_node2(x_15, x_78, x_33, x_33); +lean_inc(x_33); lean_inc(x_79); -lean_inc(x_41); +lean_inc(x_35); lean_inc(x_66); -lean_inc(x_38); -x_80 = l_Lean_Syntax_node4(x_38, x_66, x_41, x_75, x_79, x_13); -lean_inc(x_38); -x_81 = l_Lean_Syntax_node4(x_38, x_47, x_48, x_57, x_64, x_80); -lean_inc(x_29); -lean_inc(x_38); -x_82 = l_Lean_Syntax_node2(x_38, x_29, x_45, x_81); +lean_inc(x_15); +x_80 = l_Lean_Syntax_node4(x_15, x_66, x_35, x_75, x_79, x_33); +lean_inc(x_15); +x_81 = l_Lean_Syntax_node4(x_15, x_47, x_48, x_57, x_64, x_80); +lean_inc(x_36); +lean_inc(x_15); +x_82 = l_Lean_Syntax_node2(x_15, x_36, x_45, x_81); x_83 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_17); -lean_inc_ref(x_31); lean_inc_ref(x_26); -x_84 = l_Lean_Name_mkStr4(x_26, x_31, x_17, x_83); +lean_inc_ref(x_12); +lean_inc_ref(x_38); +x_84 = l_Lean_Name_mkStr4(x_38, x_12, x_26, x_83); x_85 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_38); +lean_inc(x_15); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_38); +lean_ctor_set(x_86, 0, x_15); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Syntax_SepArray_ofElems(x_28, x_30); -lean_dec_ref(x_30); -x_88 = l_Array_append___redArg(x_24, x_87); +x_87 = l_Lean_Syntax_SepArray_ofElems(x_21, x_19); +lean_dec_ref(x_19); +x_88 = l_Array_append___redArg(x_31, x_87); lean_dec_ref(x_87); -lean_inc(x_35); -lean_inc(x_38); +lean_inc(x_22); +lean_inc(x_15); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_38); -lean_ctor_set(x_89, 1, x_35); +lean_ctor_set(x_89, 0, x_15); +lean_ctor_set(x_89, 1, x_22); lean_ctor_set(x_89, 2, x_88); x_90 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_38); +lean_inc(x_15); x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_38); +lean_ctor_set(x_91, 0, x_15); lean_ctor_set(x_91, 1, x_90); -lean_inc(x_38); -x_92 = l_Lean_Syntax_node3(x_38, x_84, x_86, x_89, x_91); -lean_inc(x_35); -lean_inc(x_38); -x_93 = l_Lean_Syntax_node1(x_38, x_35, x_92); -lean_inc_n(x_13, 6); -lean_inc(x_38); -x_94 = l_Lean_Syntax_node7(x_38, x_39, x_13, x_93, x_13, x_13, x_13, x_13, x_13); +lean_inc(x_15); +x_92 = l_Lean_Syntax_node3(x_15, x_84, x_86, x_89, x_91); +lean_inc(x_22); +lean_inc(x_15); +x_93 = l_Lean_Syntax_node1(x_15, x_22, x_92); +lean_inc_n(x_33, 6); +lean_inc(x_15); +x_94 = l_Lean_Syntax_node7(x_15, x_25, x_33, x_93, x_33, x_33, x_33, x_33, x_33); x_95 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; -lean_inc_ref(x_31); -lean_inc_ref(x_26); -x_96 = l_Lean_Name_mkStr4(x_26, x_31, x_12, x_95); +lean_inc_ref(x_12); +lean_inc_ref(x_38); +x_96 = l_Lean_Name_mkStr4(x_38, x_12, x_28, x_95); x_97 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_38); +lean_inc(x_15); x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_38); +lean_ctor_set(x_98, 0, x_15); lean_ctor_set(x_98, 1, x_97); x_99 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_100 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_22); -lean_inc(x_37); -x_101 = l_Lean_addMacroScope(x_37, x_100, x_22); -lean_inc(x_20); -lean_inc(x_38); +lean_inc(x_24); +lean_inc(x_34); +x_101 = l_Lean_addMacroScope(x_34, x_100, x_24); +lean_inc(x_17); +lean_inc(x_15); x_102 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_102, 0, x_38); +lean_ctor_set(x_102, 0, x_15); lean_ctor_set(x_102, 1, x_99); lean_ctor_set(x_102, 2, x_101); -lean_ctor_set(x_102, 3, x_20); -lean_inc(x_13); -lean_inc(x_38); -x_103 = l_Lean_Syntax_node2(x_38, x_50, x_102, x_13); +lean_ctor_set(x_102, 3, x_17); +lean_inc(x_33); +lean_inc(x_15); +x_103 = l_Lean_Syntax_node2(x_15, x_50, x_102, x_33); x_104 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__5; x_105 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_106 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_22); -lean_inc(x_37); -x_107 = l_Lean_addMacroScope(x_37, x_106, x_22); -x_108 = l_Lean_Name_mkStr2(x_15, x_104); +lean_inc(x_24); +lean_inc(x_34); +x_107 = l_Lean_addMacroScope(x_34, x_106, x_24); +x_108 = l_Lean_Name_mkStr2(x_29, x_104); lean_inc(x_108); x_109 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_27); +lean_ctor_set(x_109, 1, x_37); x_110 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_110, 0, x_108); -lean_inc(x_20); +lean_inc(x_17); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_20); +lean_ctor_set(x_111, 1, x_17); x_112 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_112, 0, x_109); lean_ctor_set(x_112, 1, x_111); -lean_inc(x_38); +lean_inc(x_15); x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_38); +lean_ctor_set(x_113, 0, x_15); lean_ctor_set(x_113, 1, x_105); lean_ctor_set(x_113, 2, x_107); lean_ctor_set(x_113, 3, x_112); -lean_inc(x_38); -x_114 = l_Lean_Syntax_node2(x_38, x_61, x_34, x_113); -lean_inc(x_35); -lean_inc(x_38); -x_115 = l_Lean_Syntax_node1(x_38, x_35, x_114); -lean_inc(x_13); -lean_inc(x_38); -x_116 = l_Lean_Syntax_node2(x_38, x_59, x_13, x_115); +lean_inc(x_15); +x_114 = l_Lean_Syntax_node2(x_15, x_61, x_13, x_113); +lean_inc(x_22); +lean_inc(x_15); +x_115 = l_Lean_Syntax_node1(x_15, x_22, x_114); +lean_inc(x_33); +lean_inc(x_15); +x_116 = l_Lean_Syntax_node2(x_15, x_59, x_33, x_115); x_117 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_118 = l_Lean_Name_mkStr4(x_26, x_31, x_17, x_117); +x_118 = l_Lean_Name_mkStr4(x_38, x_12, x_26, x_117); x_119 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_38); +lean_inc(x_15); x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_38); +lean_ctor_set(x_120, 0, x_15); lean_ctor_set(x_120, 1, x_119); x_121 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_122 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_123 = l_Lean_addMacroScope(x_37, x_122, x_22); -lean_inc(x_38); +x_123 = l_Lean_addMacroScope(x_34, x_122, x_24); +lean_inc(x_15); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_38); +lean_ctor_set(x_124, 0, x_15); lean_ctor_set(x_124, 1, x_121); lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_20); -lean_inc(x_38); -x_125 = l_Lean_Syntax_node3(x_38, x_118, x_21, x_120, x_124); -lean_inc(x_13); -lean_inc(x_38); -x_126 = l_Lean_Syntax_node4(x_38, x_66, x_41, x_125, x_79, x_13); -lean_inc(x_38); -x_127 = l_Lean_Syntax_node5(x_38, x_96, x_98, x_103, x_116, x_126, x_13); -lean_inc(x_38); -x_128 = l_Lean_Syntax_node2(x_38, x_29, x_94, x_127); -x_129 = l_Lean_Syntax_node3(x_38, x_35, x_40, x_82, x_128); +lean_ctor_set(x_124, 3, x_17); +lean_inc(x_15); +x_125 = l_Lean_Syntax_node3(x_15, x_118, x_23, x_120, x_124); +lean_inc(x_33); +lean_inc(x_15); +x_126 = l_Lean_Syntax_node4(x_15, x_66, x_35, x_125, x_79, x_33); +lean_inc(x_15); +x_127 = l_Lean_Syntax_node5(x_15, x_96, x_98, x_103, x_116, x_126, x_33); +lean_inc(x_15); +x_128 = l_Lean_Syntax_node2(x_15, x_36, x_94, x_127); +x_129 = l_Lean_Syntax_node3(x_15, x_22, x_27, x_82, x_128); x_130 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_130, 0, x_129); return x_130; @@ -21598,91 +21060,91 @@ return x_130; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; x_149 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__0; x_150 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; -lean_inc_ref(x_145); -lean_inc(x_138); -lean_inc(x_140); +lean_inc_ref(x_136); +lean_inc(x_141); +lean_inc(x_134); x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_140); -lean_ctor_set(x_151, 1, x_138); -lean_ctor_set(x_151, 2, x_145); +lean_ctor_set(x_151, 0, x_134); +lean_ctor_set(x_151, 1, x_141); +lean_ctor_set(x_151, 2, x_136); x_152 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; -lean_inc(x_140); +lean_inc(x_134); x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_140); +lean_ctor_set(x_153, 0, x_134); lean_ctor_set(x_153, 1, x_152); x_154 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_140); +lean_inc(x_134); x_155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_155, 0, x_140); +lean_ctor_set(x_155, 0, x_134); lean_ctor_set(x_155, 1, x_154); x_156 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_157 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_143); +lean_inc(x_146); lean_inc(x_147); -x_158 = l_Lean_addMacroScope(x_147, x_157, x_143); +x_158 = l_Lean_addMacroScope(x_147, x_157, x_146); x_159 = lean_box(0); x_160 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__3; -lean_inc(x_141); +lean_inc(x_138); x_161 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_141); -lean_inc(x_140); +lean_ctor_set(x_161, 1, x_138); +lean_inc(x_134); x_162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_162, 0, x_140); +lean_ctor_set(x_162, 0, x_134); lean_ctor_set(x_162, 1, x_156); lean_ctor_set(x_162, 2, x_158); lean_ctor_set(x_162, 3, x_161); x_163 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_135); lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_164 = l_Lean_Name_mkStr4(x_146, x_133, x_135, x_163); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_164 = l_Lean_Name_mkStr4(x_142, x_132, x_133, x_163); x_165 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_135); lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_166 = l_Lean_Name_mkStr4(x_146, x_133, x_135, x_165); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_166 = l_Lean_Name_mkStr4(x_142, x_132, x_133, x_165); x_167 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_140); +lean_inc(x_134); x_168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_168, 0, x_140); +lean_ctor_set(x_168, 0, x_134); lean_ctor_set(x_168, 1, x_167); x_169 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_170 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_171 = lean_box(0); -lean_inc(x_143); +lean_inc(x_146); lean_inc(x_147); -x_172 = l_Lean_addMacroScope(x_147, x_171, x_143); +x_172 = l_Lean_addMacroScope(x_147, x_171, x_146); x_173 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__1; x_174 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__59; -lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_176 = l_Lean_Name_mkStr3(x_146, x_133, x_175); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_176 = l_Lean_Name_mkStr3(x_142, x_132, x_175); x_177 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_177, 0, x_176); x_178 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_146); -x_179 = l_Lean_Name_mkStr3(x_146, x_178, x_175); +lean_inc_ref(x_142); +x_179 = l_Lean_Name_mkStr3(x_142, x_178, x_175); x_180 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_180, 0, x_179); -lean_inc_ref(x_146); -x_181 = l_Lean_Name_mkStr2(x_146, x_178); +lean_inc_ref(x_142); +x_181 = l_Lean_Name_mkStr2(x_142, x_178); x_182 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_182, 0, x_181); -lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_183 = l_Lean_Name_mkStr2(x_146, x_133); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_183 = l_Lean_Name_mkStr2(x_142, x_132); x_184 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_184, 0, x_183); -lean_inc_ref(x_146); -x_185 = l_Lean_Name_mkStr1(x_146); +lean_inc_ref(x_142); +x_185 = l_Lean_Name_mkStr1(x_142); x_186 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_186, 0, x_185); -lean_inc(x_141); +lean_inc(x_138); x_187 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_141); +lean_ctor_set(x_187, 1, x_138); x_188 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_188, 0, x_184); lean_ctor_set(x_188, 1, x_187); @@ -21698,129 +21160,129 @@ lean_ctor_set(x_191, 1, x_190); x_192 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_192, 0, x_174); lean_ctor_set(x_192, 1, x_191); -lean_inc(x_140); +lean_inc(x_134); x_193 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_193, 0, x_140); +lean_ctor_set(x_193, 0, x_134); lean_ctor_set(x_193, 1, x_170); lean_ctor_set(x_193, 2, x_172); lean_ctor_set(x_193, 3, x_192); -lean_inc(x_140); -x_194 = l_Lean_Syntax_node1(x_140, x_169, x_193); -lean_inc(x_140); -x_195 = l_Lean_Syntax_node2(x_140, x_166, x_168, x_194); +lean_inc(x_134); +x_194 = l_Lean_Syntax_node1(x_134, x_169, x_193); +lean_inc(x_134); +x_195 = l_Lean_Syntax_node2(x_134, x_166, x_168, x_194); x_196 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__37; x_197 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__38; -lean_inc(x_140); +lean_inc(x_134); x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_140); +lean_ctor_set(x_198, 0, x_134); lean_ctor_set(x_198, 1, x_197); -lean_inc(x_140); -x_199 = l_Lean_Syntax_node1(x_140, x_196, x_198); +lean_inc(x_134); +x_199 = l_Lean_Syntax_node1(x_134, x_196, x_198); x_200 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -lean_inc(x_140); +lean_inc(x_134); x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_140); +lean_ctor_set(x_201, 0, x_134); lean_ctor_set(x_201, 1, x_200); -lean_inc(x_136); -lean_inc(x_138); -lean_inc(x_140); -x_202 = l_Lean_Syntax_node1(x_140, x_138, x_136); +lean_inc(x_143); +lean_inc(x_141); +lean_inc(x_134); +x_202 = l_Lean_Syntax_node1(x_134, x_141, x_143); lean_inc(x_199); -lean_inc(x_138); -lean_inc(x_140); -x_203 = l_Lean_Syntax_node3(x_140, x_138, x_199, x_201, x_202); +lean_inc(x_141); +lean_inc(x_134); +x_203 = l_Lean_Syntax_node3(x_134, x_141, x_199, x_201, x_202); x_204 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_140); +lean_inc(x_134); x_205 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_205, 0, x_140); +lean_ctor_set(x_205, 0, x_134); lean_ctor_set(x_205, 1, x_204); -lean_inc(x_140); -x_206 = l_Lean_Syntax_node3(x_140, x_164, x_195, x_203, x_205); +lean_inc(x_134); +x_206 = l_Lean_Syntax_node3(x_134, x_164, x_195, x_203, x_205); x_207 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_140); +lean_inc(x_134); x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_140); +lean_ctor_set(x_208, 0, x_134); lean_ctor_set(x_208, 1, x_207); x_209 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_135); lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_210 = l_Lean_Name_mkStr4(x_146, x_133, x_135, x_209); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_210 = l_Lean_Name_mkStr4(x_142, x_132, x_133, x_209); x_211 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_212 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_143); +lean_inc(x_146); lean_inc(x_147); -x_213 = l_Lean_addMacroScope(x_147, x_212, x_143); +x_213 = l_Lean_addMacroScope(x_147, x_212, x_146); x_214 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__8; x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_141); +lean_inc(x_138); x_216 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_141); +lean_ctor_set(x_216, 1, x_138); x_217 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_217, 0, x_214); lean_ctor_set(x_217, 1, x_216); -lean_inc(x_140); +lean_inc(x_134); x_218 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_218, 0, x_140); +lean_ctor_set(x_218, 0, x_134); lean_ctor_set(x_218, 1, x_211); lean_ctor_set(x_218, 2, x_213); lean_ctor_set(x_218, 3, x_217); -lean_inc(x_144); -lean_inc(x_138); -lean_inc(x_140); -x_219 = l_Lean_Syntax_node1(x_140, x_138, x_144); +lean_inc(x_135); +lean_inc(x_141); +lean_inc(x_134); +x_219 = l_Lean_Syntax_node1(x_134, x_141, x_135); lean_inc(x_210); -lean_inc(x_140); -x_220 = l_Lean_Syntax_node2(x_140, x_210, x_218, x_219); +lean_inc(x_134); +x_220 = l_Lean_Syntax_node2(x_134, x_210, x_218, x_219); lean_inc_ref(x_208); lean_inc_ref(x_155); -lean_inc(x_142); +lean_inc(x_144); lean_inc_ref(x_151); -lean_inc(x_140); -x_221 = l_Lean_Syntax_node8(x_140, x_150, x_151, x_153, x_142, x_155, x_162, x_206, x_208, x_220); +lean_inc(x_134); +x_221 = l_Lean_Syntax_node8(x_134, x_150, x_151, x_153, x_144, x_155, x_162, x_206, x_208, x_220); x_222 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_223 = l_Lean_Name_mkStr4(x_146, x_133, x_175, x_222); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_223 = l_Lean_Name_mkStr4(x_142, x_132, x_175, x_222); x_224 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_133); -lean_inc_ref(x_146); -x_225 = l_Lean_Name_mkStr4(x_146, x_133, x_175, x_224); +lean_inc_ref(x_132); +lean_inc_ref(x_142); +x_225 = l_Lean_Name_mkStr4(x_142, x_132, x_175, x_224); if (lean_obj_tag(x_5) == 0) { lean_object* x_226; x_226 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_12 = x_175; -x_13 = x_151; -x_14 = lean_box(0); -x_15 = x_149; -x_16 = x_173; -x_17 = x_135; -x_18 = x_136; -x_19 = x_137; -x_20 = x_141; -x_21 = x_142; -x_22 = x_143; +x_12 = x_132; +x_13 = x_155; +x_14 = x_199; +x_15 = x_134; +x_16 = lean_box(0); +x_17 = x_138; +x_18 = x_210; +x_19 = x_139; +x_20 = x_140; +x_21 = x_200; +x_22 = x_141; x_23 = x_144; -x_24 = x_145; -x_25 = x_199; -x_26 = x_146; -x_27 = x_159; -x_28 = x_200; -x_29 = x_223; -x_30 = x_132; -x_31 = x_133; -x_32 = x_210; -x_33 = x_134; -x_34 = x_155; -x_35 = x_138; -x_36 = x_139; -x_37 = x_147; -x_38 = x_140; -x_39 = x_225; -x_40 = x_221; -x_41 = x_208; +x_24 = x_146; +x_25 = x_225; +x_26 = x_133; +x_27 = x_221; +x_28 = x_175; +x_29 = x_149; +x_30 = x_135; +x_31 = x_136; +x_32 = x_137; +x_33 = x_151; +x_34 = x_147; +x_35 = x_208; +x_36 = x_223; +x_37 = x_159; +x_38 = x_142; +x_39 = x_143; +x_40 = x_145; +x_41 = x_173; x_42 = x_226; goto block_131; } @@ -21831,36 +21293,36 @@ x_227 = lean_ctor_get(x_5, 0); lean_inc(x_227); lean_dec_ref(x_5); x_228 = l_Array_mkArray1___redArg(x_227); -x_12 = x_175; -x_13 = x_151; -x_14 = lean_box(0); -x_15 = x_149; -x_16 = x_173; -x_17 = x_135; -x_18 = x_136; -x_19 = x_137; -x_20 = x_141; -x_21 = x_142; -x_22 = x_143; +x_12 = x_132; +x_13 = x_155; +x_14 = x_199; +x_15 = x_134; +x_16 = lean_box(0); +x_17 = x_138; +x_18 = x_210; +x_19 = x_139; +x_20 = x_140; +x_21 = x_200; +x_22 = x_141; x_23 = x_144; -x_24 = x_145; -x_25 = x_199; -x_26 = x_146; -x_27 = x_159; -x_28 = x_200; -x_29 = x_223; -x_30 = x_132; -x_31 = x_133; -x_32 = x_210; -x_33 = x_134; -x_34 = x_155; -x_35 = x_138; -x_36 = x_139; -x_37 = x_147; -x_38 = x_140; -x_39 = x_225; -x_40 = x_221; -x_41 = x_208; +x_24 = x_146; +x_25 = x_225; +x_26 = x_133; +x_27 = x_221; +x_28 = x_175; +x_29 = x_149; +x_30 = x_135; +x_31 = x_136; +x_32 = x_137; +x_33 = x_151; +x_34 = x_147; +x_35 = x_208; +x_36 = x_223; +x_37 = x_159; +x_38 = x_142; +x_39 = x_143; +x_40 = x_145; +x_41 = x_173; x_42 = x_228; goto block_131; } @@ -22660,44 +22122,44 @@ goto block_522; block_124: { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_17); -x_42 = l_Array_append___redArg(x_17, x_41); +lean_inc_ref(x_25); +x_42 = l_Array_append___redArg(x_25, x_41); lean_dec_ref(x_41); -lean_inc(x_20); -lean_inc(x_19); +lean_inc(x_40); +lean_inc(x_15); x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_19); -lean_ctor_set(x_43, 1, x_20); +lean_ctor_set(x_43, 0, x_15); +lean_ctor_set(x_43, 1, x_40); lean_ctor_set(x_43, 2, x_42); -lean_inc_n(x_15, 6); -lean_inc(x_32); -lean_inc(x_19); -x_44 = l_Lean_Syntax_node7(x_19, x_32, x_43, x_15, x_15, x_15, x_15, x_15, x_15); +lean_inc_n(x_33, 6); +lean_inc(x_21); +lean_inc(x_15); +x_44 = l_Lean_Syntax_node7(x_15, x_21, x_43, x_33, x_33, x_33, x_33, x_33, x_33); x_45 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__0; lean_inc_ref(x_28); -lean_inc_ref(x_16); -lean_inc_ref(x_13); -x_46 = l_Lean_Name_mkStr4(x_13, x_16, x_28, x_45); +lean_inc_ref(x_22); +lean_inc_ref(x_32); +x_46 = l_Lean_Name_mkStr4(x_32, x_22, x_28, x_45); x_47 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__1; -lean_inc(x_19); +lean_inc(x_15); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_19); +lean_ctor_set(x_48, 0, x_15); lean_ctor_set(x_48, 1, x_47); x_49 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__10; lean_inc_ref(x_28); -lean_inc_ref(x_16); -lean_inc_ref(x_13); -x_50 = l_Lean_Name_mkStr4(x_13, x_16, x_28, x_49); +lean_inc_ref(x_22); +lean_inc_ref(x_32); +x_50 = l_Lean_Name_mkStr4(x_32, x_22, x_28, x_49); x_51 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__11; x_52 = lean_box(2); -lean_inc(x_20); +lean_inc(x_40); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_20); +lean_ctor_set(x_53, 1, x_40); lean_ctor_set(x_53, 2, x_51); x_54 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; -lean_inc(x_30); -x_55 = lean_array_push(x_54, x_30); +lean_inc(x_26); +x_55 = lean_array_push(x_54, x_26); x_56 = lean_array_push(x_55, x_53); lean_inc(x_50); x_57 = lean_alloc_ctor(1, 3, 0); @@ -22705,287 +22167,287 @@ lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_50); lean_ctor_set(x_57, 2, x_56); x_58 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__13; -lean_inc_ref(x_16); -lean_inc_ref(x_13); -x_59 = l_Lean_Name_mkStr4(x_13, x_16, x_28, x_58); +lean_inc_ref(x_22); +lean_inc_ref(x_32); +x_59 = l_Lean_Name_mkStr4(x_32, x_22, x_28, x_58); x_60 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__4; x_61 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__5; -lean_inc(x_37); -lean_inc(x_38); -x_62 = l_Lean_addMacroScope(x_38, x_61, x_37); -x_63 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__6; lean_inc(x_24); +lean_inc(x_39); +x_62 = l_Lean_addMacroScope(x_39, x_61, x_24); +x_63 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__6; +lean_inc(x_18); x_64 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_24); -lean_inc(x_26); +lean_ctor_set(x_64, 1, x_18); +lean_inc(x_34); x_65 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_26); -lean_inc(x_19); +lean_ctor_set(x_65, 1, x_34); +lean_inc(x_15); x_66 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_66, 0, x_19); +lean_ctor_set(x_66, 0, x_15); lean_ctor_set(x_66, 1, x_60); lean_ctor_set(x_66, 2, x_62); lean_ctor_set(x_66, 3, x_65); -lean_inc(x_33); -lean_inc(x_27); -lean_inc(x_19); -x_67 = l_Lean_Syntax_node2(x_19, x_27, x_33, x_66); -lean_inc(x_20); -lean_inc(x_19); -x_68 = l_Lean_Syntax_node1(x_19, x_20, x_67); +lean_inc(x_14); +lean_inc(x_31); +lean_inc(x_15); +x_67 = l_Lean_Syntax_node2(x_15, x_31, x_14, x_66); +lean_inc(x_40); lean_inc(x_15); +x_68 = l_Lean_Syntax_node1(x_15, x_40, x_67); +lean_inc(x_33); lean_inc(x_59); -lean_inc(x_19); -x_69 = l_Lean_Syntax_node2(x_19, x_59, x_15, x_68); +lean_inc(x_15); +x_69 = l_Lean_Syntax_node2(x_15, x_59, x_33, x_68); x_70 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__8; x_71 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__10; -lean_inc(x_37); -lean_inc(x_38); -x_72 = l_Lean_addMacroScope(x_38, x_71, x_37); lean_inc(x_24); +lean_inc(x_39); +x_72 = l_Lean_addMacroScope(x_39, x_71, x_24); +lean_inc(x_18); x_73 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_24); -lean_inc(x_26); +lean_ctor_set(x_73, 1, x_18); +lean_inc(x_34); x_74 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_26); -lean_inc(x_19); +lean_ctor_set(x_74, 1, x_34); +lean_inc(x_15); x_75 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_75, 0, x_19); +lean_ctor_set(x_75, 0, x_15); lean_ctor_set(x_75, 1, x_70); lean_ctor_set(x_75, 2, x_72); lean_ctor_set(x_75, 3, x_74); -lean_inc(x_20); -lean_inc(x_19); -x_76 = l_Lean_Syntax_node2(x_19, x_20, x_22, x_39); -lean_inc(x_19); -x_77 = l_Lean_Syntax_node2(x_19, x_25, x_75, x_76); -lean_inc(x_15); -lean_inc(x_23); -lean_inc(x_14); lean_inc(x_40); -lean_inc(x_19); -x_78 = l_Lean_Syntax_node4(x_19, x_40, x_14, x_77, x_23, x_15); lean_inc(x_15); +x_76 = l_Lean_Syntax_node2(x_15, x_40, x_27, x_13); +lean_inc(x_15); +x_77 = l_Lean_Syntax_node2(x_15, x_37, x_75, x_76); +lean_inc(x_33); +lean_inc(x_30); +lean_inc(x_16); +lean_inc(x_29); +lean_inc(x_15); +x_78 = l_Lean_Syntax_node4(x_15, x_29, x_16, x_77, x_30, x_33); +lean_inc(x_33); lean_inc_ref(x_48); lean_inc(x_46); +lean_inc(x_15); +x_79 = l_Lean_Syntax_node5(x_15, x_46, x_48, x_57, x_69, x_78, x_33); lean_inc(x_19); -x_79 = l_Lean_Syntax_node5(x_19, x_46, x_48, x_57, x_69, x_78, x_15); -lean_inc(x_21); -lean_inc(x_19); -x_80 = l_Lean_Syntax_node2(x_19, x_21, x_44, x_79); +lean_inc(x_15); +x_80 = l_Lean_Syntax_node2(x_15, x_19, x_44, x_79); x_81 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__5; -lean_inc_ref(x_36); -lean_inc_ref(x_16); -lean_inc_ref(x_13); -x_82 = l_Lean_Name_mkStr4(x_13, x_16, x_36, x_81); +lean_inc_ref(x_35); +lean_inc_ref(x_22); +lean_inc_ref(x_32); +x_82 = l_Lean_Name_mkStr4(x_32, x_22, x_35, x_81); x_83 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__6; -lean_inc(x_19); +lean_inc(x_15); x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_19); +lean_ctor_set(x_84, 0, x_15); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Syntax_SepArray_ofElems(x_31, x_29); -lean_dec_ref(x_29); -x_86 = l_Array_append___redArg(x_17, x_85); +x_85 = l_Lean_Syntax_SepArray_ofElems(x_38, x_36); +lean_dec_ref(x_36); +x_86 = l_Array_append___redArg(x_25, x_85); lean_dec_ref(x_85); -lean_inc(x_20); -lean_inc(x_19); +lean_inc(x_40); +lean_inc(x_15); x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_19); -lean_ctor_set(x_87, 1, x_20); +lean_ctor_set(x_87, 0, x_15); +lean_ctor_set(x_87, 1, x_40); lean_ctor_set(x_87, 2, x_86); x_88 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__8; -lean_inc(x_19); +lean_inc(x_15); x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_19); +lean_ctor_set(x_89, 0, x_15); lean_ctor_set(x_89, 1, x_88); -lean_inc(x_19); -x_90 = l_Lean_Syntax_node3(x_19, x_82, x_84, x_87, x_89); -lean_inc(x_20); -lean_inc(x_19); -x_91 = l_Lean_Syntax_node1(x_19, x_20, x_90); -lean_inc_n(x_15, 6); -lean_inc(x_19); -x_92 = l_Lean_Syntax_node7(x_19, x_32, x_15, x_91, x_15, x_15, x_15, x_15, x_15); +lean_inc(x_15); +x_90 = l_Lean_Syntax_node3(x_15, x_82, x_84, x_87, x_89); +lean_inc(x_40); +lean_inc(x_15); +x_91 = l_Lean_Syntax_node1(x_15, x_40, x_90); +lean_inc_n(x_33, 6); +lean_inc(x_15); +x_92 = l_Lean_Syntax_node7(x_15, x_21, x_33, x_91, x_33, x_33, x_33, x_33, x_33); x_93 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__3; x_94 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__4; -lean_inc(x_37); -lean_inc(x_38); -x_95 = l_Lean_addMacroScope(x_38, x_94, x_37); -lean_inc(x_26); -lean_inc(x_19); +lean_inc(x_24); +lean_inc(x_39); +x_95 = l_Lean_addMacroScope(x_39, x_94, x_24); +lean_inc(x_34); +lean_inc(x_15); x_96 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_96, 0, x_19); +lean_ctor_set(x_96, 0, x_15); lean_ctor_set(x_96, 1, x_93); lean_ctor_set(x_96, 2, x_95); -lean_ctor_set(x_96, 3, x_26); +lean_ctor_set(x_96, 3, x_34); +lean_inc(x_33); lean_inc(x_15); -lean_inc(x_19); -x_97 = l_Lean_Syntax_node2(x_19, x_50, x_96, x_15); +x_97 = l_Lean_Syntax_node2(x_15, x_50, x_96, x_33); x_98 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__6; x_99 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__7; -lean_inc(x_37); -lean_inc(x_38); -x_100 = l_Lean_addMacroScope(x_38, x_99, x_37); +lean_inc(x_24); +lean_inc(x_39); +x_100 = l_Lean_addMacroScope(x_39, x_99, x_24); x_101 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__8; x_102 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_24); +lean_ctor_set(x_102, 1, x_18); x_103 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__9; -lean_inc(x_26); +lean_inc(x_34); x_104 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_26); +lean_ctor_set(x_104, 1, x_34); x_105 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_105, 0, x_102); lean_ctor_set(x_105, 1, x_104); -lean_inc(x_19); +lean_inc(x_15); x_106 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_106, 0, x_19); +lean_ctor_set(x_106, 0, x_15); lean_ctor_set(x_106, 1, x_98); lean_ctor_set(x_106, 2, x_100); lean_ctor_set(x_106, 3, x_105); -lean_inc(x_19); -x_107 = l_Lean_Syntax_node2(x_19, x_27, x_33, x_106); -lean_inc(x_20); -lean_inc(x_19); -x_108 = l_Lean_Syntax_node1(x_19, x_20, x_107); lean_inc(x_15); -lean_inc(x_19); -x_109 = l_Lean_Syntax_node2(x_19, x_59, x_15, x_108); +x_107 = l_Lean_Syntax_node2(x_15, x_31, x_14, x_106); +lean_inc(x_40); +lean_inc(x_15); +x_108 = l_Lean_Syntax_node1(x_15, x_40, x_107); +lean_inc(x_33); +lean_inc(x_15); +x_109 = l_Lean_Syntax_node2(x_15, x_59, x_33, x_108); x_110 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__10; -x_111 = l_Lean_Name_mkStr4(x_13, x_16, x_36, x_110); +x_111 = l_Lean_Name_mkStr4(x_32, x_22, x_35, x_110); x_112 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__11; -lean_inc(x_19); +lean_inc(x_15); x_113 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_113, 0, x_19); +lean_ctor_set(x_113, 0, x_15); lean_ctor_set(x_113, 1, x_112); x_114 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__13; x_115 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__14; -x_116 = l_Lean_addMacroScope(x_38, x_115, x_37); -lean_inc(x_19); +x_116 = l_Lean_addMacroScope(x_39, x_115, x_24); +lean_inc(x_15); x_117 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_117, 0, x_19); +lean_ctor_set(x_117, 0, x_15); lean_ctor_set(x_117, 1, x_114); lean_ctor_set(x_117, 2, x_116); -lean_ctor_set(x_117, 3, x_26); -lean_inc(x_19); -x_118 = l_Lean_Syntax_node3(x_19, x_111, x_30, x_113, x_117); +lean_ctor_set(x_117, 3, x_34); lean_inc(x_15); -lean_inc(x_19); -x_119 = l_Lean_Syntax_node4(x_19, x_40, x_14, x_118, x_23, x_15); -lean_inc(x_19); -x_120 = l_Lean_Syntax_node5(x_19, x_46, x_48, x_97, x_109, x_119, x_15); -lean_inc(x_19); -x_121 = l_Lean_Syntax_node2(x_19, x_21, x_92, x_120); -x_122 = l_Lean_Syntax_node4(x_19, x_20, x_34, x_35, x_80, x_121); +x_118 = l_Lean_Syntax_node3(x_15, x_111, x_26, x_113, x_117); +lean_inc(x_33); +lean_inc(x_15); +x_119 = l_Lean_Syntax_node4(x_15, x_29, x_16, x_118, x_30, x_33); +lean_inc(x_15); +x_120 = l_Lean_Syntax_node5(x_15, x_46, x_48, x_97, x_109, x_119, x_33); +lean_inc(x_15); +x_121 = l_Lean_Syntax_node2(x_15, x_19, x_92, x_120); +x_122 = l_Lean_Syntax_node4(x_15, x_40, x_17, x_20, x_80, x_121); x_123 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_18); +lean_ctor_set(x_123, 1, x_23); return x_123; } block_234: { lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_inc_ref(x_129); -x_156 = l_Array_append___redArg(x_129, x_155); +lean_inc_ref(x_138); +x_156 = l_Array_append___redArg(x_138, x_155); lean_dec_ref(x_155); -lean_inc(x_133); -lean_inc(x_132); +lean_inc(x_154); +lean_inc(x_129); x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_132); -lean_ctor_set(x_157, 1, x_133); +lean_ctor_set(x_157, 0, x_129); +lean_ctor_set(x_157, 1, x_154); lean_ctor_set(x_157, 2, x_156); -lean_inc(x_136); -lean_inc(x_126); -lean_inc(x_154); -lean_inc(x_132); -x_158 = l_Lean_Syntax_node4(x_132, x_154, x_126, x_147, x_136, x_157); -lean_inc(x_132); -x_159 = l_Lean_Syntax_node4(x_132, x_149, x_140, x_148, x_131, x_158); -lean_inc_n(x_127, 2); -lean_inc(x_132); -x_160 = l_Lean_Syntax_node4(x_132, x_135, x_127, x_127, x_142, x_159); +lean_inc(x_144); +lean_inc(x_128); +lean_inc(x_145); +lean_inc(x_129); +x_158 = l_Lean_Syntax_node4(x_129, x_145, x_128, x_135, x_144, x_157); +lean_inc(x_129); +x_159 = l_Lean_Syntax_node4(x_129, x_130, x_125, x_139, x_137, x_158); +lean_inc_n(x_148, 2); +lean_inc(x_129); +x_160 = l_Lean_Syntax_node4(x_129, x_143, x_148, x_148, x_149, x_159); x_161 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__27; x_162 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__28; -lean_inc(x_132); +lean_inc(x_129); x_163 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_163, 0, x_132); +lean_ctor_set(x_163, 0, x_129); lean_ctor_set(x_163, 1, x_162); x_164 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__30; x_165 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__31; -lean_inc(x_151); -lean_inc(x_152); -x_166 = l_Lean_addMacroScope(x_152, x_165, x_151); +lean_inc(x_134); +lean_inc(x_153); +x_166 = l_Lean_addMacroScope(x_153, x_165, x_134); x_167 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__32; -lean_inc(x_138); +lean_inc(x_131); x_168 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_168, 0, x_167); -lean_ctor_set(x_168, 1, x_138); -lean_inc(x_139); +lean_ctor_set(x_168, 1, x_131); +lean_inc(x_150); x_169 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_139); -lean_inc(x_132); +lean_ctor_set(x_169, 1, x_150); +lean_inc(x_129); x_170 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_170, 0, x_132); +lean_ctor_set(x_170, 0, x_129); lean_ctor_set(x_170, 1, x_164); lean_ctor_set(x_170, 2, x_166); lean_ctor_set(x_170, 3, x_169); x_171 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__35; -lean_inc_ref(x_150); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_172 = l_Lean_Name_mkStr4(x_125, x_128, x_150, x_171); +lean_inc_ref(x_151); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_172 = l_Lean_Name_mkStr4(x_147, x_132, x_151, x_171); x_173 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__22; -lean_inc_ref(x_150); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_174 = l_Lean_Name_mkStr4(x_125, x_128, x_150, x_173); +lean_inc_ref(x_151); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_174 = l_Lean_Name_mkStr4(x_147, x_132, x_151, x_173); x_175 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__23; -lean_inc(x_132); +lean_inc(x_129); x_176 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_176, 0, x_132); +lean_ctor_set(x_176, 0, x_129); lean_ctor_set(x_176, 1, x_175); x_177 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__25; x_178 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__27; x_179 = lean_box(0); -lean_inc(x_151); -lean_inc(x_152); -x_180 = l_Lean_addMacroScope(x_152, x_179, x_151); +lean_inc(x_134); +lean_inc(x_153); +x_180 = l_Lean_addMacroScope(x_153, x_179, x_134); x_181 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__29; -lean_inc_ref(x_143); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_182 = l_Lean_Name_mkStr3(x_125, x_128, x_143); +lean_inc_ref(x_142); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_182 = l_Lean_Name_mkStr3(x_147, x_132, x_142); x_183 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_183, 0, x_182); x_184 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__30; -lean_inc_ref(x_143); -lean_inc_ref(x_125); -x_185 = l_Lean_Name_mkStr3(x_125, x_184, x_143); +lean_inc_ref(x_142); +lean_inc_ref(x_147); +x_185 = l_Lean_Name_mkStr3(x_147, x_184, x_142); x_186 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_186, 0, x_185); -lean_inc_ref(x_125); -x_187 = l_Lean_Name_mkStr2(x_125, x_184); +lean_inc_ref(x_147); +x_187 = l_Lean_Name_mkStr2(x_147, x_184); x_188 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_188, 0, x_187); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_189 = l_Lean_Name_mkStr2(x_125, x_128); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_189 = l_Lean_Name_mkStr2(x_147, x_132); x_190 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_190, 0, x_189); -lean_inc_ref(x_125); -x_191 = l_Lean_Name_mkStr1(x_125); +lean_inc_ref(x_147); +x_191 = l_Lean_Name_mkStr1(x_147); x_192 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_192, 0, x_191); -lean_inc(x_139); +lean_inc(x_150); x_193 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_139); +lean_ctor_set(x_193, 1, x_150); x_194 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_194, 0, x_190); lean_ctor_set(x_194, 1, x_193); @@ -23001,125 +22463,125 @@ lean_ctor_set(x_197, 1, x_196); x_198 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_198, 0, x_181); lean_ctor_set(x_198, 1, x_197); -lean_inc(x_132); +lean_inc(x_129); x_199 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_199, 0, x_132); +lean_ctor_set(x_199, 0, x_129); lean_ctor_set(x_199, 1, x_178); lean_ctor_set(x_199, 2, x_180); lean_ctor_set(x_199, 3, x_198); -lean_inc(x_132); -x_200 = l_Lean_Syntax_node1(x_132, x_177, x_199); -lean_inc(x_132); -x_201 = l_Lean_Syntax_node2(x_132, x_174, x_176, x_200); +lean_inc(x_129); +x_200 = l_Lean_Syntax_node1(x_129, x_177, x_199); +lean_inc(x_129); +x_201 = l_Lean_Syntax_node2(x_129, x_174, x_176, x_200); x_202 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__37; x_203 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__38; -lean_inc(x_132); +lean_inc(x_129); x_204 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_204, 0, x_132); +lean_ctor_set(x_204, 0, x_129); lean_ctor_set(x_204, 1, x_203); -lean_inc(x_132); -x_205 = l_Lean_Syntax_node1(x_132, x_202, x_204); +lean_inc(x_129); +x_205 = l_Lean_Syntax_node1(x_129, x_202, x_204); x_206 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__7; -lean_inc(x_132); +lean_inc(x_129); x_207 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_207, 0, x_132); +lean_ctor_set(x_207, 0, x_129); lean_ctor_set(x_207, 1, x_206); -lean_inc(x_153); -lean_inc(x_133); -lean_inc(x_132); -x_208 = l_Lean_Syntax_node1(x_132, x_133, x_153); +lean_inc(x_126); +lean_inc(x_154); +lean_inc(x_129); +x_208 = l_Lean_Syntax_node1(x_129, x_154, x_126); lean_inc(x_205); -lean_inc(x_133); -lean_inc(x_132); -x_209 = l_Lean_Syntax_node3(x_132, x_133, x_205, x_207, x_208); +lean_inc(x_154); +lean_inc(x_129); +x_209 = l_Lean_Syntax_node3(x_129, x_154, x_205, x_207, x_208); x_210 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__34; -lean_inc(x_132); +lean_inc(x_129); x_211 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_211, 0, x_132); +lean_ctor_set(x_211, 0, x_129); lean_ctor_set(x_211, 1, x_210); -lean_inc(x_132); -x_212 = l_Lean_Syntax_node3(x_132, x_172, x_201, x_209, x_211); +lean_inc(x_129); +x_212 = l_Lean_Syntax_node3(x_129, x_172, x_201, x_209, x_211); x_213 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__15; -lean_inc_ref(x_150); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_214 = l_Lean_Name_mkStr4(x_125, x_128, x_150, x_213); +lean_inc_ref(x_151); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_214 = l_Lean_Name_mkStr4(x_147, x_132, x_151, x_213); x_215 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__5; x_216 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__6; -lean_inc(x_151); -lean_inc(x_152); -x_217 = l_Lean_addMacroScope(x_152, x_216, x_151); +lean_inc(x_134); +lean_inc(x_153); +x_217 = l_Lean_addMacroScope(x_153, x_216, x_134); x_218 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__7; -lean_inc(x_138); +lean_inc(x_131); x_219 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_138); +lean_ctor_set(x_219, 1, x_131); x_220 = l___private_Lake_DSL_Targets_0__Lake_DSL_mkConfigDeclDef___closed__9; -lean_inc(x_139); +lean_inc(x_150); x_221 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_139); +lean_ctor_set(x_221, 1, x_150); x_222 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_222, 0, x_219); lean_ctor_set(x_222, 1, x_221); -lean_inc(x_132); +lean_inc(x_129); x_223 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_223, 0, x_132); +lean_ctor_set(x_223, 0, x_129); lean_ctor_set(x_223, 1, x_215); lean_ctor_set(x_223, 2, x_217); lean_ctor_set(x_223, 3, x_222); -lean_inc(x_133); -lean_inc(x_132); -x_224 = l_Lean_Syntax_node1(x_132, x_133, x_137); +lean_inc(x_154); +lean_inc(x_129); +x_224 = l_Lean_Syntax_node1(x_129, x_154, x_136); lean_inc(x_214); -lean_inc(x_132); -x_225 = l_Lean_Syntax_node2(x_132, x_214, x_223, x_224); -lean_inc(x_126); -lean_inc(x_146); -lean_inc(x_145); +lean_inc(x_129); +x_225 = l_Lean_Syntax_node2(x_129, x_214, x_223, x_224); +lean_inc(x_128); lean_inc(x_127); -lean_inc(x_132); -x_226 = l_Lean_Syntax_node8(x_132, x_161, x_127, x_163, x_145, x_146, x_170, x_212, x_126, x_225); +lean_inc(x_140); +lean_inc(x_148); +lean_inc(x_129); +x_226 = l_Lean_Syntax_node8(x_129, x_161, x_148, x_163, x_140, x_127, x_170, x_212, x_128, x_225); x_227 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__40; -lean_inc_ref(x_143); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_228 = l_Lean_Name_mkStr4(x_125, x_128, x_143, x_227); +lean_inc_ref(x_142); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_228 = l_Lean_Name_mkStr4(x_147, x_132, x_142, x_227); x_229 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__41; -lean_inc_ref(x_143); -lean_inc_ref(x_128); -lean_inc_ref(x_125); -x_230 = l_Lean_Name_mkStr4(x_125, x_128, x_143, x_229); -if (lean_obj_tag(x_134) == 0) +lean_inc_ref(x_142); +lean_inc_ref(x_132); +lean_inc_ref(x_147); +x_230 = l_Lean_Name_mkStr4(x_147, x_132, x_142, x_229); +if (lean_obj_tag(x_141) == 0) { lean_object* x_231; x_231 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_13 = x_125; -x_14 = x_126; -x_15 = x_127; +x_13 = x_126; +x_14 = x_127; +x_15 = x_129; x_16 = x_128; -x_17 = x_129; -x_18 = x_130; -x_19 = x_132; -x_20 = x_133; -x_21 = x_228; -x_22 = x_205; -x_23 = x_136; -x_24 = x_138; -x_25 = x_214; -x_26 = x_139; -x_27 = x_141; -x_28 = x_143; -x_29 = x_144; -x_30 = x_145; -x_31 = x_206; -x_32 = x_230; -x_33 = x_146; -x_34 = x_160; -x_35 = x_226; -x_36 = x_150; -x_37 = x_151; -x_38 = x_152; +x_17 = x_160; +x_18 = x_131; +x_19 = x_228; +x_20 = x_226; +x_21 = x_230; +x_22 = x_132; +x_23 = x_133; +x_24 = x_134; +x_25 = x_138; +x_26 = x_140; +x_27 = x_205; +x_28 = x_142; +x_29 = x_145; +x_30 = x_144; +x_31 = x_146; +x_32 = x_147; +x_33 = x_148; +x_34 = x_150; +x_35 = x_151; +x_36 = x_152; +x_37 = x_214; +x_38 = x_206; x_39 = x_153; x_40 = x_154; x_41 = x_231; @@ -23128,36 +22590,36 @@ goto block_124; else { lean_object* x_232; lean_object* x_233; -x_232 = lean_ctor_get(x_134, 0); +x_232 = lean_ctor_get(x_141, 0); lean_inc(x_232); -lean_dec_ref(x_134); +lean_dec_ref(x_141); x_233 = l_Array_mkArray1___redArg(x_232); -x_13 = x_125; -x_14 = x_126; -x_15 = x_127; +x_13 = x_126; +x_14 = x_127; +x_15 = x_129; x_16 = x_128; -x_17 = x_129; -x_18 = x_130; -x_19 = x_132; -x_20 = x_133; -x_21 = x_228; -x_22 = x_205; -x_23 = x_136; -x_24 = x_138; -x_25 = x_214; -x_26 = x_139; -x_27 = x_141; -x_28 = x_143; -x_29 = x_144; -x_30 = x_145; -x_31 = x_206; -x_32 = x_230; -x_33 = x_146; -x_34 = x_160; -x_35 = x_226; -x_36 = x_150; -x_37 = x_151; -x_38 = x_152; +x_17 = x_160; +x_18 = x_131; +x_19 = x_228; +x_20 = x_226; +x_21 = x_230; +x_22 = x_132; +x_23 = x_133; +x_24 = x_134; +x_25 = x_138; +x_26 = x_140; +x_27 = x_205; +x_28 = x_142; +x_29 = x_145; +x_30 = x_144; +x_31 = x_146; +x_32 = x_147; +x_33 = x_148; +x_34 = x_150; +x_35 = x_151; +x_36 = x_152; +x_37 = x_214; +x_38 = x_206; x_39 = x_153; x_40 = x_154; x_41 = x_233; @@ -23167,85 +22629,85 @@ goto block_124; block_281: { lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -lean_inc_ref(x_238); -x_261 = l_Array_append___redArg(x_238, x_260); +lean_inc_ref(x_244); +x_261 = l_Array_append___redArg(x_244, x_260); lean_dec_ref(x_260); -lean_inc(x_241); -lean_inc(x_240); +lean_inc(x_259); +lean_inc(x_237); x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_240); -lean_ctor_set(x_262, 1, x_241); +lean_ctor_set(x_262, 0, x_237); +lean_ctor_set(x_262, 1, x_259); lean_ctor_set(x_262, 2, x_261); x_263 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__57; -lean_inc_ref(x_255); -lean_inc_ref(x_237); -lean_inc_ref(x_235); -x_264 = l_Lean_Name_mkStr4(x_235, x_237, x_255, x_263); +lean_inc_ref(x_254); +lean_inc_ref(x_239); +lean_inc_ref(x_251); +x_264 = l_Lean_Name_mkStr4(x_251, x_239, x_254, x_263); x_265 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__39; -lean_inc(x_240); +lean_inc(x_237); x_266 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_266, 0, x_240); +lean_ctor_set(x_266, 0, x_237); lean_ctor_set(x_266, 1, x_265); x_267 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__12; x_268 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__13; -lean_inc(x_256); -lean_inc(x_257); -x_269 = l_Lean_addMacroScope(x_257, x_268, x_256); +lean_inc(x_240); +lean_inc(x_258); +x_269 = l_Lean_addMacroScope(x_258, x_268, x_240); x_270 = lean_box(0); x_271 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___closed__16; -lean_inc(x_246); +lean_inc(x_253); x_272 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_246); -lean_inc(x_240); +lean_ctor_set(x_272, 1, x_253); +lean_inc(x_237); x_273 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_273, 0, x_240); +lean_ctor_set(x_273, 0, x_237); lean_ctor_set(x_273, 1, x_267); lean_ctor_set(x_273, 2, x_269); lean_ctor_set(x_273, 3, x_272); lean_inc_ref(x_266); lean_inc(x_264); -lean_inc(x_240); -x_274 = l_Lean_Syntax_node2(x_240, x_264, x_266, x_273); +lean_inc(x_237); +x_274 = l_Lean_Syntax_node2(x_237, x_264, x_266, x_273); x_275 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__14; -lean_inc(x_240); +lean_inc(x_237); x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_240); +lean_ctor_set(x_276, 0, x_237); lean_ctor_set(x_276, 1, x_275); -lean_inc_n(x_236, 2); -lean_inc(x_240); -x_277 = l_Lean_Syntax_node2(x_240, x_248, x_236, x_236); -if (lean_obj_tag(x_245) == 0) +lean_inc_n(x_252, 2); +lean_inc(x_237); +x_277 = l_Lean_Syntax_node2(x_237, x_247, x_252, x_252); +if (lean_obj_tag(x_256) == 0) { lean_object* x_278; x_278 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_125 = x_235; -x_126 = x_276; -x_127 = x_236; -x_128 = x_237; -x_129 = x_238; -x_130 = x_239; -x_131 = x_274; -x_132 = x_240; +x_125 = x_236; +x_126 = x_235; +x_127 = x_266; +x_128 = x_276; +x_129 = x_237; +x_130 = x_238; +x_131 = x_270; +x_132 = x_239; x_133 = x_241; -x_134 = x_242; -x_135 = x_243; -x_136 = x_277; -x_137 = x_244; -x_138 = x_270; -x_139 = x_246; -x_140 = x_247; -x_141 = x_264; -x_142 = x_249; -x_143 = x_250; -x_144 = x_251; -x_145 = x_252; -x_146 = x_266; -x_147 = x_253; -x_148 = x_262; -x_149 = x_254; -x_150 = x_255; -x_151 = x_256; +x_134 = x_240; +x_135 = x_242; +x_136 = x_243; +x_137 = x_274; +x_138 = x_244; +x_139 = x_262; +x_140 = x_245; +x_141 = x_246; +x_142 = x_248; +x_143 = x_249; +x_144 = x_277; +x_145 = x_250; +x_146 = x_264; +x_147 = x_251; +x_148 = x_252; +x_149 = x_255; +x_150 = x_253; +x_151 = x_254; x_152 = x_257; x_153 = x_258; x_154 = x_259; @@ -23255,37 +22717,37 @@ goto block_234; else { lean_object* x_279; lean_object* x_280; -x_279 = lean_ctor_get(x_245, 0); +x_279 = lean_ctor_get(x_256, 0); lean_inc(x_279); -lean_dec_ref(x_245); +lean_dec_ref(x_256); x_280 = l_Array_mkArray1___redArg(x_279); -x_125 = x_235; -x_126 = x_276; -x_127 = x_236; -x_128 = x_237; -x_129 = x_238; -x_130 = x_239; -x_131 = x_274; -x_132 = x_240; +x_125 = x_236; +x_126 = x_235; +x_127 = x_266; +x_128 = x_276; +x_129 = x_237; +x_130 = x_238; +x_131 = x_270; +x_132 = x_239; x_133 = x_241; -x_134 = x_242; -x_135 = x_243; -x_136 = x_277; -x_137 = x_244; -x_138 = x_270; -x_139 = x_246; -x_140 = x_247; -x_141 = x_264; -x_142 = x_249; -x_143 = x_250; -x_144 = x_251; -x_145 = x_252; -x_146 = x_266; -x_147 = x_253; -x_148 = x_262; -x_149 = x_254; -x_150 = x_255; -x_151 = x_256; +x_134 = x_240; +x_135 = x_242; +x_136 = x_243; +x_137 = x_274; +x_138 = x_244; +x_139 = x_262; +x_140 = x_245; +x_141 = x_246; +x_142 = x_248; +x_143 = x_249; +x_144 = x_277; +x_145 = x_250; +x_146 = x_264; +x_147 = x_251; +x_148 = x_252; +x_149 = x_255; +x_150 = x_253; +x_151 = x_254; x_152 = x_257; x_153 = x_258; x_154 = x_259; @@ -23297,92 +22759,92 @@ goto block_234; { uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; x_306 = 0; -x_307 = l_Lean_mkIdentFrom(x_292, x_305, x_306); -lean_inc(x_292); -x_308 = l_Lake_Name_quoteFrom(x_292, x_290, x_306); +x_307 = l_Lean_mkIdentFrom(x_286, x_305, x_306); +lean_inc(x_286); +x_308 = l_Lake_Name_quoteFrom(x_286, x_290, x_306); x_309 = l_Lake_ExternLib_keyword; x_310 = l_Lake_Name_quoteFrom(x_282, x_309, x_306); -x_311 = l_Lean_SourceInfo_fromRef(x_294, x_306); -lean_dec(x_294); +x_311 = l_Lean_SourceInfo_fromRef(x_292, x_306); +lean_dec(x_292); x_312 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__16; -lean_inc_ref(x_293); -lean_inc(x_299); +lean_inc_ref(x_283); +lean_inc(x_304); lean_inc(x_311); x_313 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_313, 0, x_311); -lean_ctor_set(x_313, 1, x_299); -lean_ctor_set(x_313, 2, x_293); +lean_ctor_set(x_313, 1, x_304); +lean_ctor_set(x_313, 2, x_283); lean_inc(x_311); x_314 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_314, 0, x_311); -lean_ctor_set(x_314, 1, x_283); +lean_ctor_set(x_314, 1, x_297); x_315 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__68; lean_inc(x_311); -x_316 = l_Lean_Syntax_node1(x_311, x_295, x_307); -if (lean_obj_tag(x_288) == 0) +x_316 = l_Lean_Syntax_node1(x_311, x_284, x_307); +if (lean_obj_tag(x_285) == 0) { lean_object* x_317; x_317 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__35; -x_235 = x_284; -x_236 = x_313; -x_237 = x_286; -x_238 = x_293; -x_239 = x_297; -x_240 = x_311; -x_241 = x_299; -x_242 = x_301; -x_243 = x_312; -x_244 = x_310; -x_245 = x_304; -x_246 = x_285; -x_247 = x_316; -x_248 = x_287; -x_249 = x_314; -x_250 = x_289; -x_251 = x_291; -x_252 = x_292; +x_235 = x_308; +x_236 = x_316; +x_237 = x_311; +x_238 = x_315; +x_239 = x_298; +x_240 = x_300; +x_241 = x_301; +x_242 = x_303; +x_243 = x_310; +x_244 = x_283; +x_245 = x_286; +x_246 = x_287; +x_247 = x_289; +x_248 = x_288; +x_249 = x_312; +x_250 = x_291; +x_251 = x_293; +x_252 = x_313; x_253 = x_296; -x_254 = x_315; -x_255 = x_298; -x_256 = x_300; -x_257 = x_302; -x_258 = x_308; -x_259 = x_303; +x_254 = x_295; +x_255 = x_314; +x_256 = x_294; +x_257 = x_299; +x_258 = x_302; +x_259 = x_304; x_260 = x_317; goto block_281; } else { lean_object* x_318; lean_object* x_319; -x_318 = lean_ctor_get(x_288, 0); +x_318 = lean_ctor_get(x_285, 0); lean_inc(x_318); -lean_dec_ref(x_288); +lean_dec_ref(x_285); x_319 = l_Array_mkArray1___redArg(x_318); -x_235 = x_284; -x_236 = x_313; -x_237 = x_286; -x_238 = x_293; -x_239 = x_297; -x_240 = x_311; -x_241 = x_299; -x_242 = x_301; -x_243 = x_312; -x_244 = x_310; -x_245 = x_304; -x_246 = x_285; -x_247 = x_316; -x_248 = x_287; -x_249 = x_314; -x_250 = x_289; -x_251 = x_291; -x_252 = x_292; +x_235 = x_308; +x_236 = x_316; +x_237 = x_311; +x_238 = x_315; +x_239 = x_298; +x_240 = x_300; +x_241 = x_301; +x_242 = x_303; +x_243 = x_310; +x_244 = x_283; +x_245 = x_286; +x_246 = x_287; +x_247 = x_289; +x_248 = x_288; +x_249 = x_312; +x_250 = x_291; +x_251 = x_293; +x_252 = x_313; x_253 = x_296; -x_254 = x_315; -x_255 = x_298; -x_256 = x_300; -x_257 = x_302; -x_258 = x_308; -x_259 = x_303; +x_254 = x_295; +x_255 = x_314; +x_256 = x_294; +x_257 = x_299; +x_258 = x_302; +x_259 = x_304; x_260 = x_319; goto block_281; } @@ -23419,13 +22881,13 @@ lean_dec_ref(x_342); x_345 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__41; x_346 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__56; x_347 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_348 = l_Lean_Name_mkStr4(x_321, x_323, x_346, x_347); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_348 = l_Lean_Name_mkStr4(x_330, x_322, x_346, x_347); x_349 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_350 = l_Lean_Name_mkStr4(x_321, x_323, x_346, x_349); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_350 = l_Lean_Name_mkStr4(x_330, x_322, x_346, x_349); x_351 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_352 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_340); @@ -23439,9 +22901,9 @@ lean_inc(x_340); x_354 = l_Lean_Syntax_node1(x_340, x_350, x_353); x_355 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_356 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_357 = l_Lean_Name_mkStr4(x_321, x_323, x_355, x_356); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_357 = l_Lean_Name_mkStr4(x_330, x_322, x_355, x_356); x_358 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_359 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__24; x_360 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; @@ -23485,10 +22947,10 @@ x_373 = l_Lean_Syntax_node2(x_343, x_348, x_367, x_372); x_374 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; x_375 = lean_array_push(x_374, x_365); x_376 = lean_array_push(x_375, x_373); -x_377 = l_Lake_DSL_expandAttrs(x_325); +x_377 = l_Lake_DSL_expandAttrs(x_321); x_378 = l_Array_append___redArg(x_376, x_377); lean_dec_ref(x_377); -x_379 = l_Lake_DSL_expandIdentOrStrAsIdent(x_322); +x_379 = l_Lake_DSL_expandIdentOrStrAsIdent(x_329); x_380 = l_Lean_TSyntax_getId(x_379); x_381 = l_Lean_Name_hasMacroScopes(x_380); if (x_381 == 0) @@ -23496,28 +22958,28 @@ if (x_381 == 0) lean_object* x_382; lean_inc(x_380); x_382 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___lam__1(x_380); -x_283 = x_359; -x_284 = x_321; -x_285 = x_362; -x_286 = x_323; -x_287 = x_326; -x_288 = x_328; -x_289 = x_327; +x_283 = x_352; +x_284 = x_345; +x_285 = x_323; +x_286 = x_379; +x_287 = x_325; +x_288 = x_327; +x_289 = x_326; x_290 = x_380; -x_291 = x_378; -x_292 = x_379; -x_293 = x_352; -x_294 = x_338; -x_295 = x_345; -x_296 = x_330; -x_297 = x_344; -x_298 = x_346; -x_299 = x_351; +x_291 = x_328; +x_292 = x_338; +x_293 = x_330; +x_294 = x_331; +x_295 = x_346; +x_296 = x_362; +x_297 = x_359; +x_298 = x_322; +x_299 = x_378; x_300 = x_336; -x_301 = x_324; +x_301 = x_344; x_302 = x_335; -x_303 = x_329; -x_304 = x_331; +x_303 = x_324; +x_304 = x_351; x_305 = x_382; goto block_320; } @@ -23534,28 +22996,28 @@ x_385 = lean_ctor_get(x_383, 0); x_386 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___lam__1(x_385); lean_ctor_set(x_383, 0, x_386); x_387 = l_Lean_MacroScopesView_review(x_383); -x_283 = x_359; -x_284 = x_321; -x_285 = x_362; -x_286 = x_323; -x_287 = x_326; -x_288 = x_328; -x_289 = x_327; +x_283 = x_352; +x_284 = x_345; +x_285 = x_323; +x_286 = x_379; +x_287 = x_325; +x_288 = x_327; +x_289 = x_326; x_290 = x_380; -x_291 = x_378; -x_292 = x_379; -x_293 = x_352; -x_294 = x_338; -x_295 = x_345; -x_296 = x_330; -x_297 = x_344; -x_298 = x_346; -x_299 = x_351; +x_291 = x_328; +x_292 = x_338; +x_293 = x_330; +x_294 = x_331; +x_295 = x_346; +x_296 = x_362; +x_297 = x_359; +x_298 = x_322; +x_299 = x_378; x_300 = x_336; -x_301 = x_324; +x_301 = x_344; x_302 = x_335; -x_303 = x_329; -x_304 = x_331; +x_303 = x_324; +x_304 = x_351; x_305 = x_387; goto block_320; } @@ -23578,28 +23040,28 @@ lean_ctor_set(x_393, 1, x_389); lean_ctor_set(x_393, 2, x_390); lean_ctor_set(x_393, 3, x_391); x_394 = l_Lean_MacroScopesView_review(x_393); -x_283 = x_359; -x_284 = x_321; -x_285 = x_362; -x_286 = x_323; -x_287 = x_326; -x_288 = x_328; -x_289 = x_327; +x_283 = x_352; +x_284 = x_345; +x_285 = x_323; +x_286 = x_379; +x_287 = x_325; +x_288 = x_327; +x_289 = x_326; x_290 = x_380; -x_291 = x_378; -x_292 = x_379; -x_293 = x_352; -x_294 = x_338; -x_295 = x_345; -x_296 = x_330; -x_297 = x_344; -x_298 = x_346; -x_299 = x_351; +x_291 = x_328; +x_292 = x_338; +x_293 = x_330; +x_294 = x_331; +x_295 = x_346; +x_296 = x_362; +x_297 = x_359; +x_298 = x_322; +x_299 = x_378; x_300 = x_336; -x_301 = x_324; +x_301 = x_344; x_302 = x_335; -x_303 = x_329; -x_304 = x_331; +x_303 = x_324; +x_304 = x_351; x_305 = x_394; goto block_320; } @@ -23649,13 +23111,13 @@ lean_dec_ref(x_406); x_409 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__41; x_410 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__56; x_411 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__42; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_412 = l_Lean_Name_mkStr4(x_321, x_323, x_410, x_411); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_412 = l_Lean_Name_mkStr4(x_330, x_322, x_410, x_411); x_413 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__43; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_414 = l_Lean_Name_mkStr4(x_321, x_323, x_410, x_413); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_414 = l_Lean_Name_mkStr4(x_330, x_322, x_410, x_413); x_415 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__45; x_416 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__46; lean_inc(x_404); @@ -23669,9 +23131,9 @@ lean_inc(x_404); x_418 = l_Lean_Syntax_node1(x_404, x_414, x_417); x_419 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__47; x_420 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__48; -lean_inc_ref(x_323); -lean_inc_ref(x_321); -x_421 = l_Lean_Name_mkStr4(x_321, x_323, x_419, x_420); +lean_inc_ref(x_322); +lean_inc_ref(x_330); +x_421 = l_Lean_Name_mkStr4(x_330, x_322, x_419, x_420); x_422 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__23; x_423 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__24; x_424 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandTargetCommand___closed__25; @@ -23715,10 +23177,10 @@ x_437 = l_Lean_Syntax_node2(x_407, x_412, x_431, x_436); x_438 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandModuleFacetDecl___closed__12; x_439 = lean_array_push(x_438, x_429); x_440 = lean_array_push(x_439, x_437); -x_441 = l_Lake_DSL_expandAttrs(x_325); +x_441 = l_Lake_DSL_expandAttrs(x_321); x_442 = l_Array_append___redArg(x_440, x_441); lean_dec_ref(x_441); -x_443 = l_Lake_DSL_expandIdentOrStrAsIdent(x_322); +x_443 = l_Lake_DSL_expandIdentOrStrAsIdent(x_329); x_444 = l_Lean_TSyntax_getId(x_443); x_445 = l_Lean_Name_hasMacroScopes(x_444); if (x_445 == 0) @@ -23726,28 +23188,28 @@ if (x_445 == 0) lean_object* x_446; lean_inc(x_444); x_446 = l___private_Lake_DSL_Targets_0__Lake_DSL_expandExternLibCommand___lam__1(x_444); -x_283 = x_423; -x_284 = x_321; -x_285 = x_426; -x_286 = x_323; -x_287 = x_326; -x_288 = x_328; -x_289 = x_327; +x_283 = x_416; +x_284 = x_409; +x_285 = x_323; +x_286 = x_443; +x_287 = x_325; +x_288 = x_327; +x_289 = x_326; x_290 = x_444; -x_291 = x_442; -x_292 = x_443; -x_293 = x_416; -x_294 = x_401; -x_295 = x_409; -x_296 = x_330; -x_297 = x_408; -x_298 = x_410; -x_299 = x_415; +x_291 = x_328; +x_292 = x_401; +x_293 = x_330; +x_294 = x_331; +x_295 = x_410; +x_296 = x_426; +x_297 = x_423; +x_298 = x_322; +x_299 = x_442; x_300 = x_397; -x_301 = x_324; +x_301 = x_408; x_302 = x_396; -x_303 = x_329; -x_304 = x_331; +x_303 = x_324; +x_304 = x_415; x_305 = x_446; goto block_320; } @@ -23785,28 +23247,28 @@ lean_ctor_set(x_454, 1, x_449); lean_ctor_set(x_454, 2, x_450); lean_ctor_set(x_454, 3, x_451); x_455 = l_Lean_MacroScopesView_review(x_454); -x_283 = x_423; -x_284 = x_321; -x_285 = x_426; -x_286 = x_323; -x_287 = x_326; -x_288 = x_328; -x_289 = x_327; +x_283 = x_416; +x_284 = x_409; +x_285 = x_323; +x_286 = x_443; +x_287 = x_325; +x_288 = x_327; +x_289 = x_326; x_290 = x_444; -x_291 = x_442; -x_292 = x_443; -x_293 = x_416; -x_294 = x_401; -x_295 = x_409; -x_296 = x_330; -x_297 = x_408; -x_298 = x_410; -x_299 = x_415; +x_291 = x_328; +x_292 = x_401; +x_293 = x_330; +x_294 = x_331; +x_295 = x_410; +x_296 = x_426; +x_297 = x_423; +x_298 = x_322; +x_299 = x_442; x_300 = x_397; -x_301 = x_324; +x_301 = x_408; x_302 = x_396; -x_303 = x_329; -x_304 = x_331; +x_303 = x_324; +x_304 = x_415; x_305 = x_455; goto block_320; } @@ -23954,16 +23416,16 @@ lean_object* x_497; lean_dec(x_458); x_497 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_497, 0, x_492); -x_321 = x_466; -x_322 = x_459; -x_323 = x_467; -x_324 = x_461; +x_321 = x_459; +x_322 = x_467; +x_323 = x_462; +x_324 = x_486; x_325 = x_460; x_326 = x_474; x_327 = x_468; -x_328 = x_462; -x_329 = x_469; -x_330 = x_486; +x_328 = x_469; +x_329 = x_461; +x_330 = x_466; x_331 = x_497; x_332 = x_463; x_333 = x_464; @@ -23977,16 +23439,16 @@ lean_object* x_498; lean_dec(x_487); lean_dec(x_458); x_498 = lean_box(0); -x_321 = x_466; -x_322 = x_459; -x_323 = x_467; -x_324 = x_461; +x_321 = x_459; +x_322 = x_467; +x_323 = x_462; +x_324 = x_486; x_325 = x_460; x_326 = x_474; x_327 = x_468; -x_328 = x_462; -x_329 = x_469; -x_330 = x_486; +x_328 = x_469; +x_329 = x_461; +x_330 = x_466; x_331 = x_498; x_332 = x_463; x_333 = x_464; @@ -24044,9 +23506,9 @@ x_512 = l_Lean_Syntax_getArg(x_507, x_8); lean_dec(x_507); x_513 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_513, 0, x_512); -x_459 = x_506; -x_460 = x_501; -x_461 = x_502; +x_459 = x_501; +x_460 = x_502; +x_461 = x_506; x_462 = x_513; x_463 = x_2; x_464 = x_3; @@ -24058,9 +23520,9 @@ else lean_object* x_514; lean_dec(x_507); x_514 = lean_box(0); -x_459 = x_506; -x_460 = x_501; -x_461 = x_502; +x_459 = x_501; +x_460 = x_502; +x_461 = x_506; x_462 = x_514; x_463 = x_2; x_464 = x_3; diff --git a/stage0/stdlib/Lake/Load/Toml.c b/stage0/stdlib/Lake/Load/Toml.c index 7ed962e14d14..8b13ec632ab5 100644 --- a/stage0/stdlib/Lake/Load/Toml.c +++ b/stage0/stdlib/Lake/Load/Toml.c @@ -63425,9 +63425,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 { @@ -63446,7 +63446,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; } @@ -63472,8 +63472,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; } @@ -63481,8 +63481,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; } @@ -65288,7 +65288,7 @@ lean_inc_ref(x_20); x_21 = l_Lake_Toml_loadToml(x_20); if (lean_obj_tag(x_21) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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; size_t x_62; lean_object* x_63; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; size_t x_84; lean_object* x_85; lean_object* x_86; 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_118; lean_object* x_119; lean_object* x_120; lean_object* x_144; lean_object* x_145; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_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_118; lean_object* x_119; lean_object* x_120; lean_object* x_144; lean_object* x_145; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec_ref(x_21); @@ -65343,35 +65343,35 @@ goto block_148; block_52: { uint8_t x_38; -x_38 = l_Array_isEmpty___redArg(x_24); +x_38 = l_Array_isEmpty___redArg(x_29); if (x_38 == 0) { lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_dec_ref(x_37); -lean_dec_ref(x_35); +lean_dec_ref(x_36); lean_dec(x_34); -lean_dec_ref(x_33); -lean_dec(x_32); +lean_dec(x_33); lean_dec_ref(x_31); lean_dec_ref(x_30); -lean_dec_ref(x_29); -lean_dec(x_27); +lean_dec_ref(x_28); +lean_dec_ref(x_27); lean_dec_ref(x_26); -lean_dec(x_25); -lean_dec_ref(x_23); +lean_dec_ref(x_25); +lean_dec(x_24); +lean_dec(x_23); 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); -x_39 = lean_array_get_size(x_24); +x_39 = lean_array_get_size(x_29); x_40 = lean_array_get_size(x_2); -x_41 = lean_nat_dec_lt(x_36, x_39); +x_41 = lean_nat_dec_lt(x_35, x_39); if (x_41 == 0) { lean_dec(x_39); -lean_dec_ref(x_24); +lean_dec_ref(x_29); lean_dec_ref(x_20); x_4 = x_40; x_5 = x_2; @@ -65385,7 +65385,7 @@ x_42 = lean_nat_dec_le(x_39, x_39); if (x_42 == 0) { lean_dec(x_39); -lean_dec_ref(x_24); +lean_dec_ref(x_29); lean_dec_ref(x_20); x_4 = x_40; x_5 = x_2; @@ -65398,8 +65398,8 @@ lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; x_43 = lean_box(0); x_44 = lean_usize_of_nat(x_39); lean_dec(x_39); -x_45 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_20, x_38, x_24, x_28, x_44, x_43, x_2); -lean_dec_ref(x_24); +x_45 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lake_loadTomlConfig_spec__1(x_20, x_38, x_29, x_32, x_44, x_43, x_2); +lean_dec_ref(x_29); x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); lean_dec_ref(x_45); @@ -65413,31 +65413,31 @@ goto block_8; else { lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec_ref(x_24); +lean_dec_ref(x_29); lean_dec_ref(x_20); -x_47 = lean_ctor_get(x_30, 13); +x_47 = lean_ctor_get(x_25, 13); lean_inc_ref(x_47); -x_48 = lean_ctor_get(x_30, 15); +x_48 = lean_ctor_get(x_25, 15); lean_inc_ref(x_48); x_49 = lean_box(0); x_50 = lean_alloc_ctor(0, 22, 0); -lean_ctor_set(x_50, 0, x_27); -lean_ctor_set(x_50, 1, x_25); +lean_ctor_set(x_50, 0, x_34); +lean_ctor_set(x_50, 1, x_23); lean_ctor_set(x_50, 2, x_11); lean_ctor_set(x_50, 3, x_10); -lean_ctor_set(x_50, 4, x_30); +lean_ctor_set(x_50, 4, x_25); lean_ctor_set(x_50, 5, x_13); lean_ctor_set(x_50, 6, x_12); -lean_ctor_set(x_50, 7, x_23); +lean_ctor_set(x_50, 7, x_28); lean_ctor_set(x_50, 8, x_14); lean_ctor_set(x_50, 9, x_15); -lean_ctor_set(x_50, 10, x_33); +lean_ctor_set(x_50, 10, x_36); lean_ctor_set(x_50, 11, x_26); -lean_ctor_set(x_50, 12, x_34); -lean_ctor_set(x_50, 13, x_31); -lean_ctor_set(x_50, 14, x_32); -lean_ctor_set(x_50, 15, x_29); -lean_ctor_set(x_50, 16, x_35); +lean_ctor_set(x_50, 12, x_33); +lean_ctor_set(x_50, 13, x_27); +lean_ctor_set(x_50, 14, x_24); +lean_ctor_set(x_50, 15, x_30); +lean_ctor_set(x_50, 16, x_31); lean_ctor_set(x_50, 17, x_37); lean_ctor_set(x_50, 18, x_47); lean_ctor_set(x_50, 19, x_48); @@ -65456,7 +65456,7 @@ x_64 = l_System_FilePath_normalize(x_63); x_65 = lean_box(1); x_66 = lean_unsigned_to_nat(0u); x_67 = l_Lake_loadTomlConfig___closed__0; -if (lean_obj_tag(x_55) == 0) +if (lean_obj_tag(x_58) == 0) { uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; x_68 = 0; @@ -65468,43 +65468,43 @@ x_72 = l_Lake_loadTomlConfig___closed__1; x_73 = lean_string_append(x_71, x_72); x_74 = l_Lake_loadTomlConfig___closed__2; x_75 = lean_string_append(x_73, x_74); -x_23 = x_64; -x_24 = x_57; -x_25 = x_56; -x_26 = x_60; -x_27 = x_61; -x_28 = x_62; -x_29 = x_67; -x_30 = x_54; -x_31 = x_53; -x_32 = x_65; -x_33 = x_59; -x_34 = x_58; -x_35 = x_67; -x_36 = x_66; +x_23 = x_54; +x_24 = x_65; +x_25 = x_57; +x_26 = x_59; +x_27 = x_60; +x_28 = x_64; +x_29 = x_53; +x_30 = x_67; +x_31 = x_67; +x_32 = x_55; +x_33 = x_56; +x_34 = x_61; +x_35 = x_66; +x_36 = x_62; x_37 = x_75; goto block_52; } else { lean_object* x_76; -x_76 = lean_ctor_get(x_55, 0); +x_76 = lean_ctor_get(x_58, 0); lean_inc(x_76); -lean_dec_ref(x_55); -x_23 = x_64; -x_24 = x_57; -x_25 = x_56; -x_26 = x_60; -x_27 = x_61; -x_28 = x_62; -x_29 = x_67; -x_30 = x_54; -x_31 = x_53; -x_32 = x_65; -x_33 = x_59; -x_34 = x_58; -x_35 = x_67; -x_36 = x_66; +lean_dec_ref(x_58); +x_23 = x_54; +x_24 = x_65; +x_25 = x_57; +x_26 = x_59; +x_27 = x_60; +x_28 = x_64; +x_29 = x_53; +x_30 = x_67; +x_31 = x_67; +x_32 = x_55; +x_33 = x_56; +x_34 = x_61; +x_35 = x_66; +x_36 = x_62; x_37 = x_76; goto block_52; } @@ -65512,43 +65512,43 @@ goto block_52; block_92: { lean_object* x_87; -x_87 = lean_ctor_get(x_78, 2); +x_87 = lean_ctor_get(x_81, 2); if (lean_obj_tag(x_87) == 0) { lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_78, 12); +x_88 = lean_ctor_get(x_81, 12); lean_inc(x_88); x_89 = l_Lake_loadTomlConfig___closed__3; -x_53 = x_79; +x_53 = x_86; x_54 = x_78; -x_55 = x_88; +x_55 = x_79; x_56 = x_80; -x_57 = x_86; -x_58 = x_82; -x_59 = x_85; -x_60 = x_81; -x_61 = x_83; -x_62 = x_84; +x_57 = x_81; +x_58 = x_88; +x_59 = x_82; +x_60 = x_83; +x_61 = x_84; +x_62 = x_85; x_63 = x_89; goto block_77; } else { lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_78, 12); +x_90 = lean_ctor_get(x_81, 12); lean_inc(x_90); x_91 = lean_ctor_get(x_87, 0); lean_inc(x_91); -x_53 = x_79; +x_53 = x_86; x_54 = x_78; -x_55 = x_90; +x_55 = x_79; x_56 = x_80; -x_57 = x_86; -x_58 = x_82; -x_59 = x_85; -x_60 = x_81; -x_61 = x_83; -x_62 = x_84; +x_57 = x_81; +x_58 = x_90; +x_59 = x_82; +x_60 = x_83; +x_61 = x_84; +x_62 = x_85; x_63 = x_91; goto block_77; } @@ -65566,12 +65566,12 @@ x_106 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_105, x_103, x_22); if (lean_obj_tag(x_106) == 0) { x_78 = x_93; -x_79 = x_102; +x_79 = x_101; x_80 = x_94; -x_81 = x_96; -x_82 = x_95; -x_83 = x_97; -x_84 = x_101; +x_81 = x_95; +x_82 = x_96; +x_83 = x_102; +x_84 = x_97; x_85 = x_104; x_86 = x_99; goto block_92; @@ -65605,12 +65605,12 @@ x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec_ref(x_112); x_78 = x_93; -x_79 = x_102; +x_79 = x_101; x_80 = x_94; -x_81 = x_96; -x_82 = x_95; -x_83 = x_97; -x_84 = x_101; +x_81 = x_95; +x_82 = x_96; +x_83 = x_102; +x_84 = x_97; x_85 = x_113; x_86 = x_114; goto block_92; @@ -65622,12 +65622,12 @@ x_115 = lean_ctor_get(x_112, 1); lean_inc(x_115); lean_dec_ref(x_112); x_78 = x_93; -x_79 = x_102; +x_79 = x_101; x_80 = x_94; -x_81 = x_96; -x_82 = x_95; -x_83 = x_97; -x_84 = x_101; +x_81 = x_95; +x_82 = x_96; +x_83 = x_102; +x_84 = x_97; x_85 = x_104; x_86 = x_115; goto block_92; @@ -65640,12 +65640,12 @@ x_116 = lean_ctor_get(x_109, 1); lean_inc(x_116); lean_dec_ref(x_109); x_78 = x_93; -x_79 = x_102; +x_79 = x_101; x_80 = x_94; -x_81 = x_96; -x_82 = x_95; -x_83 = x_97; -x_84 = x_101; +x_81 = x_95; +x_82 = x_96; +x_83 = x_102; +x_84 = x_97; x_85 = x_104; x_86 = x_116; goto block_92; @@ -65656,9 +65656,9 @@ goto block_92; { 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_inc(x_22); -lean_inc(x_119); +lean_inc(x_118); lean_inc(x_120); -x_121 = l___private_Lake_Load_Toml_0__Lake_decodeTomlConfig___at___00Lake_PackageConfig_decodeToml_spec__0(x_120, x_119, x_22, x_118); +x_121 = l___private_Lake_Load_Toml_0__Lake_decodeTomlConfig___at___00Lake_PackageConfig_decodeToml_spec__0(x_120, x_118, x_22, x_119); x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); x_123 = lean_ctor_get(x_121, 1); @@ -65684,9 +65684,9 @@ lean_inc(x_22); x_132 = l_Lake_Toml_RBDict_findEntry_x3f___redArg(x_131, x_129, x_22); if (lean_obj_tag(x_132) == 0) { -x_93 = x_122; -x_94 = x_119; -x_95 = x_128; +x_93 = x_118; +x_94 = x_128; +x_95 = x_122; x_96 = x_127; x_97 = x_120; x_98 = x_130; @@ -65721,9 +65721,9 @@ lean_inc(x_139); x_140 = lean_ctor_get(x_138, 1); lean_inc(x_140); lean_dec_ref(x_138); -x_93 = x_122; -x_94 = x_119; -x_95 = x_128; +x_93 = x_118; +x_94 = x_128; +x_95 = x_122; x_96 = x_127; x_97 = x_120; x_98 = x_139; @@ -65736,9 +65736,9 @@ lean_object* x_141; x_141 = lean_ctor_get(x_138, 1); lean_inc(x_141); lean_dec_ref(x_138); -x_93 = x_122; -x_94 = x_119; -x_95 = x_128; +x_93 = x_118; +x_94 = x_128; +x_95 = x_122; x_96 = x_127; x_97 = x_120; x_98 = x_130; @@ -65752,9 +65752,9 @@ lean_object* x_142; x_142 = lean_ctor_get(x_135, 1); lean_inc(x_142); lean_dec_ref(x_135); -x_93 = x_122; -x_94 = x_119; -x_95 = x_128; +x_93 = x_118; +x_94 = x_128; +x_95 = x_122; x_96 = x_127; x_97 = x_120; x_98 = x_130; @@ -65770,8 +65770,8 @@ x_146 = l_Lake_stringToLegalOrSimpleName(x_144); x_147 = l_Lean_Name_isAnonymous(x_9); if (x_147 == 0) { -x_118 = x_145; -x_119 = x_146; +x_118 = x_146; +x_119 = x_145; x_120 = x_9; goto block_143; } @@ -65779,8 +65779,8 @@ else { lean_dec(x_9); lean_inc(x_146); -x_118 = x_145; -x_119 = x_146; +x_118 = x_146; +x_119 = x_145; x_120 = x_146; goto block_143; } diff --git a/stage0/stdlib/Lake/Util/OpaqueType.c b/stage0/stdlib/Lake/Util/OpaqueType.c index 035047338cc6..236d91d5d13d 100644 --- a/stage0/stdlib/Lake/Util/OpaqueType.c +++ b/stage0/stdlib/Lake/Util/OpaqueType.c @@ -3244,7 +3244,7 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; 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; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; 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_226; lean_object* x_336; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; lean_object* x_56; lean_object* x_57; size_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_226; lean_object* x_336; x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(2u); @@ -3289,452 +3289,452 @@ goto block_335; block_225: { lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; size_t x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -lean_inc_ref(x_22); -x_69 = l_Array_append___redArg(x_22, x_68); +lean_inc_ref(x_65); +x_69 = l_Array_append___redArg(x_65, x_68); lean_dec_ref(x_68); -lean_inc(x_54); lean_inc(x_29); +lean_inc(x_22); x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_29); -lean_ctor_set(x_70, 1, x_54); +lean_ctor_set(x_70, 0, x_22); +lean_ctor_set(x_70, 1, x_29); lean_ctor_set(x_70, 2, x_69); lean_inc_ref(x_70); -lean_inc_n(x_63, 5); -lean_inc(x_32); -lean_inc(x_29); -x_71 = l_Lean_Syntax_node7(x_29, x_32, x_63, x_21, x_70, x_63, x_63, x_63, x_63); +lean_inc_n(x_47, 5); +lean_inc(x_44); +lean_inc(x_22); +x_71 = l_Lean_Syntax_node7(x_22, x_44, x_47, x_36, x_70, x_47, x_47, x_47, x_47); x_72 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__nonemptyTypeCmd__1___closed__42; -lean_inc_ref(x_33); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_73 = l_Lean_Name_mkStr4(x_64, x_60, x_33, x_72); -lean_inc(x_29); +lean_inc_ref(x_30); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_73 = l_Lean_Name_mkStr4(x_39, x_53, x_30, x_72); +lean_inc(x_22); x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_29); +lean_ctor_set(x_74, 0, x_22); lean_ctor_set(x_74, 1, x_72); -lean_inc(x_46); -lean_inc_ref(x_35); -x_75 = lean_array_push(x_35, x_46); -lean_inc(x_48); -x_76 = lean_array_push(x_75, x_48); -lean_inc(x_17); -lean_inc(x_39); +lean_inc(x_67); +lean_inc_ref(x_24); +x_75 = lean_array_push(x_24, x_67); +lean_inc(x_37); +x_76 = lean_array_push(x_75, x_37); +lean_inc(x_32); +lean_inc(x_56); x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_39); -lean_ctor_set(x_77, 1, x_17); +lean_ctor_set(x_77, 0, x_56); +lean_ctor_set(x_77, 1, x_32); lean_ctor_set(x_77, 2, x_76); x_78 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__nonemptyTypeCmd__1___closed__49; -lean_inc_ref(x_33); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_79 = l_Lean_Name_mkStr4(x_64, x_60, x_33, x_78); -lean_inc(x_63); +lean_inc_ref(x_30); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_79 = l_Lean_Name_mkStr4(x_39, x_53, x_30, x_78); +lean_inc(x_47); lean_inc(x_79); -lean_inc(x_29); -x_80 = l_Lean_Syntax_node2(x_29, x_79, x_63, x_40); -lean_inc(x_63); +lean_inc(x_22); +x_80 = l_Lean_Syntax_node2(x_22, x_79, x_47, x_38); +lean_inc(x_47); lean_inc_ref(x_74); lean_inc(x_73); -lean_inc(x_29); -x_81 = l_Lean_Syntax_node4(x_29, x_73, x_74, x_77, x_80, x_63); -lean_inc(x_55); -lean_inc(x_29); -x_82 = l_Lean_Syntax_node2(x_29, x_55, x_71, x_81); +lean_inc(x_22); +x_81 = l_Lean_Syntax_node4(x_22, x_73, x_74, x_77, x_80, x_47); +lean_inc(x_49); +lean_inc(x_22); +x_82 = l_Lean_Syntax_node2(x_22, x_49, x_71, x_81); lean_inc_ref(x_70); -lean_inc_n(x_63, 6); -lean_inc(x_32); -lean_inc(x_29); -x_83 = l_Lean_Syntax_node7(x_29, x_32, x_63, x_63, x_70, x_63, x_63, x_63, x_63); +lean_inc_n(x_47, 6); +lean_inc(x_44); +lean_inc(x_22); +x_83 = l_Lean_Syntax_node7(x_22, x_44, x_47, x_47, x_70, x_47, x_47, x_47, x_47); x_84 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__nonemptyTypeCmd__1___closed__13; -lean_inc_ref(x_33); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_85 = l_Lean_Name_mkStr4(x_64, x_60, x_33, x_84); -lean_inc(x_29); +lean_inc_ref(x_30); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_85 = l_Lean_Name_mkStr4(x_39, x_53, x_30, x_84); +lean_inc(x_22); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_29); +lean_ctor_set(x_86, 0, x_22); lean_ctor_set(x_86, 1, x_84); -lean_inc(x_63); -lean_inc(x_17); -lean_inc(x_29); -x_87 = l_Lean_Syntax_node2(x_29, x_17, x_49, x_63); -lean_inc(x_54); +lean_inc(x_47); +lean_inc(x_32); +lean_inc(x_22); +x_87 = l_Lean_Syntax_node2(x_22, x_32, x_33, x_47); lean_inc(x_29); -x_88 = l_Lean_Syntax_node1(x_29, x_54, x_87); +lean_inc(x_22); +x_88 = l_Lean_Syntax_node1(x_22, x_29, x_87); x_89 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__1; x_90 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__2; -lean_inc(x_38); -lean_inc(x_66); -x_91 = l_Lean_addMacroScope(x_66, x_90, x_38); -lean_inc(x_36); +lean_inc(x_60); +lean_inc(x_64); +x_91 = l_Lean_addMacroScope(x_64, x_90, x_60); +lean_inc(x_48); x_92 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_36); +lean_ctor_set(x_92, 1, x_48); x_93 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__3; -lean_inc(x_53); +lean_inc(x_52); x_94 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_53); +lean_ctor_set(x_94, 1, x_52); x_95 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_95, 0, x_92); lean_ctor_set(x_95, 1, x_94); -lean_inc(x_29); +lean_inc(x_22); x_96 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_96, 0, x_29); +lean_ctor_set(x_96, 0, x_22); lean_ctor_set(x_96, 1, x_89); lean_ctor_set(x_96, 2, x_91); lean_ctor_set(x_96, 3, x_95); x_97 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__4; -lean_inc_ref(x_34); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_98 = l_Lean_Name_mkStr4(x_64, x_60, x_34, x_97); +lean_inc_ref(x_19); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_98 = l_Lean_Name_mkStr4(x_39, x_53, x_19, x_97); x_99 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__5; -lean_inc_ref(x_34); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_100 = l_Lean_Name_mkStr4(x_64, x_60, x_34, x_99); +lean_inc_ref(x_19); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_100 = l_Lean_Name_mkStr4(x_39, x_53, x_19, x_99); x_101 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__6; -lean_inc(x_29); +lean_inc(x_22); x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_29); +lean_ctor_set(x_102, 0, x_22); lean_ctor_set(x_102, 1, x_101); x_103 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__8; x_104 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__10; x_105 = lean_box(0); -lean_inc(x_38); -lean_inc(x_66); -x_106 = l_Lean_addMacroScope(x_66, x_105, x_38); +lean_inc(x_60); +lean_inc(x_64); +x_106 = l_Lean_addMacroScope(x_64, x_105, x_60); x_107 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__12; -lean_inc_ref(x_64); -x_108 = l_Lean_Name_mkStr1(x_64); +lean_inc_ref(x_39); +x_108 = l_Lean_Name_mkStr1(x_39); x_109 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_109, 0, x_108); -lean_inc(x_53); +lean_inc(x_52); x_110 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_53); +lean_ctor_set(x_110, 1, x_52); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_107); lean_ctor_set(x_111, 1, x_110); -lean_inc(x_29); +lean_inc(x_22); x_112 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_112, 0, x_29); +lean_ctor_set(x_112, 0, x_22); lean_ctor_set(x_112, 1, x_104); lean_ctor_set(x_112, 2, x_106); lean_ctor_set(x_112, 3, x_111); -lean_inc(x_29); -x_113 = l_Lean_Syntax_node1(x_29, x_103, x_112); -lean_inc(x_29); -x_114 = l_Lean_Syntax_node2(x_29, x_100, x_102, x_113); +lean_inc(x_22); +x_113 = l_Lean_Syntax_node1(x_22, x_103, x_112); +lean_inc(x_22); +x_114 = l_Lean_Syntax_node2(x_22, x_100, x_102, x_113); x_115 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__13; -lean_inc(x_29); +lean_inc(x_22); x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_29); +lean_ctor_set(x_116, 0, x_22); lean_ctor_set(x_116, 1, x_115); lean_inc_ref(x_116); -lean_inc(x_23); +lean_inc(x_46); lean_inc(x_114); lean_inc(x_98); -lean_inc(x_29); -x_117 = l_Lean_Syntax_node3(x_29, x_98, x_114, x_23, x_116); -lean_inc(x_37); -lean_inc(x_29); -x_118 = l_Lean_Syntax_node3(x_29, x_98, x_114, x_37, x_116); +lean_inc(x_22); +x_117 = l_Lean_Syntax_node3(x_22, x_98, x_114, x_46, x_116); +lean_inc(x_51); +lean_inc(x_22); +x_118 = l_Lean_Syntax_node3(x_22, x_98, x_114, x_51, x_116); lean_inc(x_118); lean_inc(x_117); -lean_inc(x_54); lean_inc(x_29); -x_119 = l_Lean_Syntax_node2(x_29, x_54, x_117, x_118); +lean_inc(x_22); +x_119 = l_Lean_Syntax_node2(x_22, x_29, x_117, x_118); lean_inc_ref(x_96); -lean_inc(x_58); -lean_inc(x_29); -x_120 = l_Lean_Syntax_node2(x_29, x_58, x_96, x_119); lean_inc(x_41); -lean_inc(x_19); -lean_inc(x_29); -x_121 = l_Lean_Syntax_node2(x_29, x_19, x_41, x_120); -lean_inc(x_63); +lean_inc(x_22); +x_120 = l_Lean_Syntax_node2(x_22, x_41, x_96, x_119); +lean_inc(x_28); +lean_inc(x_45); +lean_inc(x_22); +x_121 = l_Lean_Syntax_node2(x_22, x_45, x_28, x_120); +lean_inc(x_47); lean_inc(x_79); -lean_inc(x_29); -x_122 = l_Lean_Syntax_node2(x_29, x_79, x_63, x_121); +lean_inc(x_22); +x_122 = l_Lean_Syntax_node2(x_22, x_79, x_47, x_121); x_123 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__14; -lean_inc_ref(x_34); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_124 = l_Lean_Name_mkStr4(x_64, x_60, x_34, x_123); +lean_inc_ref(x_19); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_124 = l_Lean_Name_mkStr4(x_39, x_53, x_19, x_123); x_125 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__15; -lean_inc(x_29); +lean_inc(x_22); x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_29); +lean_ctor_set(x_126, 0, x_22); lean_ctor_set(x_126, 1, x_125); -lean_inc(x_46); -lean_inc(x_54); +lean_inc(x_67); lean_inc(x_29); -x_127 = l_Lean_Syntax_node1(x_29, x_54, x_46); +lean_inc(x_22); +x_127 = l_Lean_Syntax_node1(x_22, x_29, x_67); x_128 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__16; -lean_inc(x_29); +lean_inc(x_22); x_129 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_129, 0, x_29); +lean_ctor_set(x_129, 0, x_22); lean_ctor_set(x_129, 1, x_128); lean_inc_ref(x_129); lean_inc_ref(x_126); lean_inc(x_124); -lean_inc(x_29); -x_130 = l_Lean_Syntax_node3(x_29, x_124, x_126, x_127, x_129); -lean_inc(x_63); +lean_inc(x_22); +x_130 = l_Lean_Syntax_node3(x_22, x_124, x_126, x_127, x_129); +lean_inc(x_47); +lean_inc(x_62); +lean_inc(x_61); lean_inc(x_59); -lean_inc(x_26); -lean_inc(x_57); -lean_inc(x_29); -x_131 = l_Lean_Syntax_node4(x_29, x_57, x_26, x_130, x_59, x_63); -lean_inc(x_63); +lean_inc(x_22); +x_131 = l_Lean_Syntax_node4(x_22, x_59, x_61, x_130, x_62, x_47); +lean_inc(x_47); lean_inc_ref(x_86); -lean_inc(x_51); +lean_inc(x_34); lean_inc(x_85); -lean_inc(x_29); -x_132 = l_Lean_Syntax_node6(x_29, x_85, x_51, x_86, x_63, x_88, x_122, x_131); +lean_inc(x_22); +x_132 = l_Lean_Syntax_node6(x_22, x_85, x_34, x_86, x_47, x_88, x_122, x_131); lean_inc(x_83); -lean_inc(x_55); -lean_inc(x_29); -x_133 = l_Lean_Syntax_node2(x_29, x_55, x_83, x_132); -lean_inc(x_61); -lean_inc_ref(x_35); -x_134 = lean_array_push(x_35, x_61); -lean_inc(x_48); -x_135 = lean_array_push(x_134, x_48); -lean_inc(x_17); -lean_inc(x_39); +lean_inc(x_49); +lean_inc(x_22); +x_133 = l_Lean_Syntax_node2(x_22, x_49, x_83, x_132); +lean_inc(x_35); +lean_inc_ref(x_24); +x_134 = lean_array_push(x_24, x_35); +lean_inc(x_37); +x_135 = lean_array_push(x_134, x_37); +lean_inc(x_32); +lean_inc(x_56); x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_39); -lean_ctor_set(x_136, 1, x_17); +lean_ctor_set(x_136, 0, x_56); +lean_ctor_set(x_136, 1, x_32); lean_ctor_set(x_136, 2, x_135); -lean_inc(x_29); -x_137 = l_Lean_Syntax_node3(x_29, x_31, x_37, x_42, x_23); -lean_inc(x_41); -lean_inc(x_19); -lean_inc(x_29); -x_138 = l_Lean_Syntax_node2(x_29, x_19, x_41, x_137); +lean_inc(x_22); +x_137 = l_Lean_Syntax_node3(x_22, x_31, x_51, x_23, x_46); +lean_inc(x_28); +lean_inc(x_45); +lean_inc(x_22); +x_138 = l_Lean_Syntax_node2(x_22, x_45, x_28, x_137); lean_inc(x_138); -lean_inc(x_54); -lean_inc(x_29); -x_139 = l_Lean_Syntax_node1(x_29, x_54, x_138); -lean_inc(x_63); lean_inc(x_29); -x_140 = l_Lean_Syntax_node2(x_29, x_24, x_63, x_139); -lean_inc(x_63); +lean_inc(x_22); +x_139 = l_Lean_Syntax_node1(x_22, x_29, x_138); +lean_inc(x_47); +lean_inc(x_22); +x_140 = l_Lean_Syntax_node2(x_22, x_20, x_47, x_139); +lean_inc(x_47); +lean_inc(x_22); +x_141 = l_Lean_Syntax_node5(x_22, x_26, x_66, x_136, x_140, x_25, x_47); +lean_inc(x_49); +lean_inc(x_22); +x_142 = l_Lean_Syntax_node2(x_22, x_49, x_54, x_141); lean_inc(x_29); -x_141 = l_Lean_Syntax_node5(x_29, x_56, x_65, x_136, x_140, x_28, x_63); -lean_inc(x_55); -lean_inc(x_29); -x_142 = l_Lean_Syntax_node2(x_29, x_55, x_18, x_141); -lean_inc(x_54); -lean_inc(x_29); -x_143 = l_Lean_Syntax_node1(x_29, x_54, x_61); -lean_inc(x_29); -x_144 = l_Lean_Syntax_node2(x_29, x_47, x_27, x_143); -lean_inc(x_51); -lean_inc(x_29); -x_145 = l_Lean_Syntax_node2(x_29, x_50, x_51, x_144); -lean_inc(x_54); -lean_inc(x_29); -x_146 = l_Lean_Syntax_node1(x_29, x_54, x_145); -lean_inc(x_67); +lean_inc(x_22); +x_143 = l_Lean_Syntax_node1(x_22, x_29, x_35); +lean_inc(x_22); +x_144 = l_Lean_Syntax_node2(x_22, x_40, x_27, x_143); +lean_inc(x_34); +lean_inc(x_22); +x_145 = l_Lean_Syntax_node2(x_22, x_43, x_34, x_144); lean_inc(x_29); -x_147 = l_Lean_Syntax_node3(x_29, x_45, x_20, x_146, x_67); -lean_inc(x_54); +lean_inc(x_22); +x_146 = l_Lean_Syntax_node1(x_22, x_29, x_145); +lean_inc(x_18); +lean_inc(x_22); +x_147 = l_Lean_Syntax_node3(x_22, x_63, x_57, x_146, x_18); lean_inc(x_29); -x_148 = l_Lean_Syntax_node1(x_29, x_54, x_147); -lean_inc_n(x_63, 5); -lean_inc(x_29); -x_149 = l_Lean_Syntax_node7(x_29, x_32, x_63, x_148, x_70, x_63, x_63, x_63, x_63); -lean_inc(x_25); -x_150 = lean_array_push(x_35, x_25); -x_151 = lean_array_push(x_150, x_48); -lean_inc(x_17); +lean_inc(x_22); +x_148 = l_Lean_Syntax_node1(x_22, x_29, x_147); +lean_inc_n(x_47, 5); +lean_inc(x_22); +x_149 = l_Lean_Syntax_node7(x_22, x_44, x_47, x_148, x_70, x_47, x_47, x_47, x_47); +lean_inc(x_21); +x_150 = lean_array_push(x_24, x_21); +x_151 = lean_array_push(x_150, x_37); +lean_inc(x_32); x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_39); -lean_ctor_set(x_152, 1, x_17); +lean_ctor_set(x_152, 0, x_56); +lean_ctor_set(x_152, 1, x_32); lean_ctor_set(x_152, 2, x_151); -x_153 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1_spec__1___redArg(x_43, x_30, x_16); +x_153 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1_spec__1___redArg(x_55, x_58, x_16); x_154 = lean_array_size(x_153); -lean_inc(x_63); -lean_inc(x_54); +lean_inc(x_47); lean_inc(x_29); -lean_inc_ref(x_34); -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_155 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1_spec__2(x_64, x_60, x_34, x_29, x_54, x_63, x_154, x_30, x_153); -x_156 = l_Array_append___redArg(x_22, x_155); +lean_inc(x_22); +lean_inc_ref(x_19); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_155 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1_spec__2(x_39, x_53, x_19, x_22, x_29, x_47, x_154, x_58, x_153); +x_156 = l_Array_append___redArg(x_65, x_155); lean_dec_ref(x_155); -lean_inc(x_54); lean_inc(x_29); +lean_inc(x_22); x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_29); -lean_ctor_set(x_157, 1, x_54); +lean_ctor_set(x_157, 0, x_22); +lean_ctor_set(x_157, 1, x_29); lean_ctor_set(x_157, 2, x_156); lean_inc(x_79); +lean_inc(x_22); +x_158 = l_Lean_Syntax_node2(x_22, x_79, x_157, x_138); +lean_inc(x_47); +lean_inc(x_22); +x_159 = l_Lean_Syntax_node4(x_22, x_73, x_74, x_152, x_158, x_47); +lean_inc(x_49); +lean_inc(x_22); +x_160 = l_Lean_Syntax_node2(x_22, x_49, x_149, x_159); +lean_inc(x_47); +lean_inc(x_22); +x_161 = l_Lean_Syntax_node2(x_22, x_32, x_50, x_47); lean_inc(x_29); -x_158 = l_Lean_Syntax_node2(x_29, x_79, x_157, x_138); -lean_inc(x_63); -lean_inc(x_29); -x_159 = l_Lean_Syntax_node4(x_29, x_73, x_74, x_152, x_158, x_63); -lean_inc(x_55); -lean_inc(x_29); -x_160 = l_Lean_Syntax_node2(x_29, x_55, x_149, x_159); -lean_inc(x_63); -lean_inc(x_29); -x_161 = l_Lean_Syntax_node2(x_29, x_17, x_62, x_63); -lean_inc(x_54); -lean_inc(x_29); -x_162 = l_Lean_Syntax_node1(x_29, x_54, x_161); +lean_inc(x_22); +x_162 = l_Lean_Syntax_node1(x_22, x_29, x_161); lean_inc(x_117); lean_inc(x_118); -lean_inc(x_54); -lean_inc(x_29); -x_163 = l_Lean_Syntax_node2(x_29, x_54, x_118, x_117); -lean_inc(x_58); lean_inc(x_29); -x_164 = l_Lean_Syntax_node2(x_29, x_58, x_96, x_163); +lean_inc(x_22); +x_163 = l_Lean_Syntax_node2(x_22, x_29, x_118, x_117); lean_inc(x_41); -lean_inc(x_19); -lean_inc(x_29); -x_165 = l_Lean_Syntax_node2(x_29, x_19, x_41, x_164); -lean_inc(x_63); +lean_inc(x_22); +x_164 = l_Lean_Syntax_node2(x_22, x_41, x_96, x_163); +lean_inc(x_28); +lean_inc(x_45); +lean_inc(x_22); +x_165 = l_Lean_Syntax_node2(x_22, x_45, x_28, x_164); +lean_inc(x_47); lean_inc(x_79); +lean_inc(x_22); +x_166 = l_Lean_Syntax_node2(x_22, x_79, x_47, x_165); lean_inc(x_29); -x_166 = l_Lean_Syntax_node2(x_29, x_79, x_63, x_165); -lean_inc(x_54); -lean_inc(x_29); -x_167 = l_Lean_Syntax_node1(x_29, x_54, x_25); +lean_inc(x_22); +x_167 = l_Lean_Syntax_node1(x_22, x_29, x_21); lean_inc_ref(x_129); lean_inc_ref(x_126); lean_inc(x_124); -lean_inc(x_29); -x_168 = l_Lean_Syntax_node3(x_29, x_124, x_126, x_167, x_129); -lean_inc(x_63); +lean_inc(x_22); +x_168 = l_Lean_Syntax_node3(x_22, x_124, x_126, x_167, x_129); +lean_inc(x_47); +lean_inc(x_62); +lean_inc(x_61); lean_inc(x_59); -lean_inc(x_26); -lean_inc(x_57); -lean_inc(x_29); -x_169 = l_Lean_Syntax_node4(x_29, x_57, x_26, x_168, x_59, x_63); -lean_inc(x_63); +lean_inc(x_22); +x_169 = l_Lean_Syntax_node4(x_22, x_59, x_61, x_168, x_62, x_47); +lean_inc(x_47); lean_inc_ref(x_86); -lean_inc(x_51); +lean_inc(x_34); lean_inc(x_85); -lean_inc(x_29); -x_170 = l_Lean_Syntax_node6(x_29, x_85, x_51, x_86, x_63, x_162, x_166, x_169); +lean_inc(x_22); +x_170 = l_Lean_Syntax_node6(x_22, x_85, x_34, x_86, x_47, x_162, x_166, x_169); lean_inc(x_83); -lean_inc(x_55); -lean_inc(x_29); -x_171 = l_Lean_Syntax_node2(x_29, x_55, x_83, x_170); +lean_inc(x_49); +lean_inc(x_22); +x_171 = l_Lean_Syntax_node2(x_22, x_49, x_83, x_170); x_172 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__17; -lean_inc_ref(x_60); -lean_inc_ref(x_64); -x_173 = l_Lean_Name_mkStr4(x_64, x_60, x_34, x_172); +lean_inc_ref(x_53); +lean_inc_ref(x_39); +x_173 = l_Lean_Name_mkStr4(x_39, x_53, x_19, x_172); x_174 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__18; -lean_inc(x_29); +lean_inc(x_22); x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_29); +lean_ctor_set(x_175, 0, x_22); lean_ctor_set(x_175, 1, x_174); x_176 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__20; x_177 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__21; -lean_inc(x_38); -lean_inc(x_66); -x_178 = l_Lean_addMacroScope(x_66, x_177, x_38); -lean_inc(x_36); +lean_inc(x_60); +lean_inc(x_64); +x_178 = l_Lean_addMacroScope(x_64, x_177, x_60); +lean_inc(x_48); x_179 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_36); +lean_ctor_set(x_179, 1, x_48); x_180 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__22; -lean_inc(x_53); +lean_inc(x_52); x_181 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_53); +lean_ctor_set(x_181, 1, x_52); x_182 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_182, 0, x_179); lean_ctor_set(x_182, 1, x_181); -lean_inc(x_29); +lean_inc(x_22); x_183 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_183, 0, x_29); +lean_ctor_set(x_183, 0, x_22); lean_ctor_set(x_183, 1, x_176); lean_ctor_set(x_183, 2, x_178); lean_ctor_set(x_183, 3, x_182); -lean_inc(x_54); lean_inc(x_29); -x_184 = l_Lean_Syntax_node1(x_29, x_54, x_117); +lean_inc(x_22); +x_184 = l_Lean_Syntax_node1(x_22, x_29, x_117); lean_inc_ref(x_183); -lean_inc(x_58); -lean_inc(x_29); -x_185 = l_Lean_Syntax_node2(x_29, x_58, x_183, x_184); -lean_inc(x_63); -lean_inc(x_29); -x_186 = l_Lean_Syntax_node4(x_29, x_173, x_175, x_63, x_185, x_67); -lean_inc(x_54); -lean_inc(x_29); -x_187 = l_Lean_Syntax_node1(x_29, x_54, x_186); -lean_inc(x_54); -lean_inc(x_29); -x_188 = l_Lean_Syntax_node1(x_29, x_54, x_118); -lean_inc(x_58); -lean_inc(x_29); -x_189 = l_Lean_Syntax_node2(x_29, x_58, x_183, x_188); +lean_inc(x_41); +lean_inc(x_22); +x_185 = l_Lean_Syntax_node2(x_22, x_41, x_183, x_184); +lean_inc(x_47); +lean_inc(x_22); +x_186 = l_Lean_Syntax_node4(x_22, x_173, x_175, x_47, x_185, x_18); lean_inc(x_29); -x_190 = l_Lean_Syntax_node2(x_29, x_19, x_41, x_189); +lean_inc(x_22); +x_187 = l_Lean_Syntax_node1(x_22, x_29, x_186); lean_inc(x_29); -x_191 = l_Lean_Syntax_node2(x_29, x_79, x_187, x_190); +lean_inc(x_22); +x_188 = l_Lean_Syntax_node1(x_22, x_29, x_118); +lean_inc(x_41); +lean_inc(x_22); +x_189 = l_Lean_Syntax_node2(x_22, x_41, x_183, x_188); +lean_inc(x_22); +x_190 = l_Lean_Syntax_node2(x_22, x_45, x_28, x_189); +lean_inc(x_22); +x_191 = l_Lean_Syntax_node2(x_22, x_79, x_187, x_190); x_192 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__24; x_193 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__25; -x_194 = l_Lean_addMacroScope(x_66, x_193, x_38); +x_194 = l_Lean_addMacroScope(x_64, x_193, x_60); x_195 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__26; x_196 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_196, 0, x_195); -lean_ctor_set(x_196, 1, x_36); +lean_ctor_set(x_196, 1, x_48); x_197 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__27; x_198 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_53); +lean_ctor_set(x_198, 1, x_52); x_199 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_199, 0, x_196); lean_ctor_set(x_199, 1, x_198); -lean_inc(x_29); +lean_inc(x_22); x_200 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_200, 0, x_29); +lean_ctor_set(x_200, 0, x_22); lean_ctor_set(x_200, 1, x_192); lean_ctor_set(x_200, 2, x_194); lean_ctor_set(x_200, 3, x_199); -lean_inc(x_54); lean_inc(x_29); -x_201 = l_Lean_Syntax_node1(x_29, x_54, x_200); +lean_inc(x_22); +x_201 = l_Lean_Syntax_node1(x_22, x_29, x_200); +lean_inc(x_22); +x_202 = l_Lean_Syntax_node2(x_22, x_41, x_67, x_201); lean_inc(x_29); -x_202 = l_Lean_Syntax_node2(x_29, x_58, x_46, x_201); -lean_inc(x_54); -lean_inc(x_29); -x_203 = l_Lean_Syntax_node1(x_29, x_54, x_202); -lean_inc(x_29); -x_204 = l_Lean_Syntax_node3(x_29, x_124, x_126, x_203, x_129); -lean_inc(x_63); -lean_inc(x_29); -x_205 = l_Lean_Syntax_node4(x_29, x_57, x_26, x_204, x_59, x_63); -lean_inc_n(x_63, 2); -lean_inc(x_29); -x_206 = l_Lean_Syntax_node6(x_29, x_85, x_51, x_86, x_63, x_63, x_191, x_205); -lean_inc(x_29); -x_207 = l_Lean_Syntax_node2(x_29, x_55, x_83, x_206); +lean_inc(x_22); +x_203 = l_Lean_Syntax_node1(x_22, x_29, x_202); +lean_inc(x_22); +x_204 = l_Lean_Syntax_node3(x_22, x_124, x_126, x_203, x_129); +lean_inc(x_47); +lean_inc(x_22); +x_205 = l_Lean_Syntax_node4(x_22, x_59, x_61, x_204, x_62, x_47); +lean_inc_n(x_47, 2); +lean_inc(x_22); +x_206 = l_Lean_Syntax_node6(x_22, x_85, x_34, x_86, x_47, x_47, x_191, x_205); +lean_inc(x_22); +x_207 = l_Lean_Syntax_node2(x_22, x_49, x_83, x_206); x_208 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__28; -x_209 = l_Lean_Name_mkStr4(x_64, x_60, x_33, x_208); -lean_inc(x_29); +x_209 = l_Lean_Name_mkStr4(x_39, x_53, x_30, x_208); +lean_inc(x_22); x_210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_210, 0, x_29); +lean_ctor_set(x_210, 0, x_22); lean_ctor_set(x_210, 1, x_208); -lean_inc(x_54); -lean_inc(x_29); -x_211 = l_Lean_Syntax_node2(x_29, x_54, x_11, x_63); lean_inc(x_29); -x_212 = l_Lean_Syntax_node2(x_29, x_209, x_210, x_211); +lean_inc(x_22); +x_211 = l_Lean_Syntax_node2(x_22, x_29, x_11, x_47); +lean_inc(x_22); +x_212 = l_Lean_Syntax_node2(x_22, x_209, x_210, x_211); x_213 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__hydrateOpaqueTypeCmd__1___closed__29; -x_214 = lean_array_push(x_213, x_44); -x_215 = lean_array_push(x_214, x_52); +x_214 = lean_array_push(x_213, x_17); +x_215 = lean_array_push(x_214, x_42); x_216 = lean_array_push(x_215, x_82); x_217 = lean_array_push(x_216, x_133); x_218 = lean_array_push(x_217, x_142); @@ -3743,8 +3743,8 @@ x_220 = lean_array_push(x_219, x_171); x_221 = lean_array_push(x_220, x_207); x_222 = lean_array_push(x_221, x_212); x_223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_223, 0, x_29); -lean_ctor_set(x_223, 1, x_54); +lean_ctor_set(x_223, 0, x_22); +lean_ctor_set(x_223, 1, x_29); lean_ctor_set(x_223, 2, x_222); x_224 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_224, 0, x_223); @@ -3981,57 +3981,57 @@ if (lean_obj_tag(x_226) == 0) { lean_object* x_332; x_332 = l_Lake___aux__Lake__Util__OpaqueType______macroRules__Lake__nonemptyTypeCmd__1___closed__29; -x_17 = x_285; -x_18 = x_281; -x_19 = x_291; -x_20 = x_253; -x_21 = x_331; -x_22 = x_248; -x_23 = x_301; -x_24 = x_290; -x_25 = x_233; -x_26 = x_311; +x_17 = x_245; +x_18 = x_268; +x_19 = x_250; +x_20 = x_290; +x_21 = x_233; +x_22 = x_237; +x_23 = x_303; +x_24 = x_288; +x_25 = x_319; +x_26 = x_282; x_27 = x_325; -x_28 = x_319; -x_29 = x_237; -x_30 = x_297; +x_28 = x_293; +x_29 = x_238; +x_30 = x_241; x_31 = x_294; -x_32 = x_247; -x_33 = x_241; -x_34 = x_250; -x_35 = x_288; -x_36 = x_261; -x_37 = x_304; -x_38 = x_228; -x_39 = x_286; -x_40 = x_306; -x_41 = x_293; -x_42 = x_303; -x_43 = x_296; -x_44 = x_245; -x_45 = x_251; -x_46 = x_230; -x_47 = x_257; -x_48 = x_287; -x_49 = x_232; -x_50 = x_254; -x_51 = x_256; -x_52 = x_321; -x_53 = x_261; -x_54 = x_238; -x_55 = x_246; -x_56 = x_282; -x_57 = x_309; -x_58 = x_295; -x_59 = x_318; -x_60 = x_240; -x_61 = x_234; -x_62 = x_235; -x_63 = x_249; -x_64 = x_239; -x_65 = x_284; -x_66 = x_227; -x_67 = x_268; +x_32 = x_285; +x_33 = x_232; +x_34 = x_256; +x_35 = x_234; +x_36 = x_331; +x_37 = x_287; +x_38 = x_306; +x_39 = x_239; +x_40 = x_257; +x_41 = x_295; +x_42 = x_321; +x_43 = x_254; +x_44 = x_247; +x_45 = x_291; +x_46 = x_301; +x_47 = x_249; +x_48 = x_261; +x_49 = x_246; +x_50 = x_235; +x_51 = x_304; +x_52 = x_261; +x_53 = x_240; +x_54 = x_281; +x_55 = x_296; +x_56 = x_286; +x_57 = x_253; +x_58 = x_297; +x_59 = x_309; +x_60 = x_228; +x_61 = x_311; +x_62 = x_318; +x_63 = x_251; +x_64 = x_227; +x_65 = x_248; +x_66 = x_284; +x_67 = x_230; x_68 = x_332; goto block_225; } @@ -4042,57 +4042,57 @@ x_333 = lean_ctor_get(x_226, 0); lean_inc(x_333); lean_dec_ref(x_226); x_334 = l_Array_mkArray1___redArg(x_333); -x_17 = x_285; -x_18 = x_281; -x_19 = x_291; -x_20 = x_253; -x_21 = x_331; -x_22 = x_248; -x_23 = x_301; -x_24 = x_290; -x_25 = x_233; -x_26 = x_311; +x_17 = x_245; +x_18 = x_268; +x_19 = x_250; +x_20 = x_290; +x_21 = x_233; +x_22 = x_237; +x_23 = x_303; +x_24 = x_288; +x_25 = x_319; +x_26 = x_282; x_27 = x_325; -x_28 = x_319; -x_29 = x_237; -x_30 = x_297; +x_28 = x_293; +x_29 = x_238; +x_30 = x_241; x_31 = x_294; -x_32 = x_247; -x_33 = x_241; -x_34 = x_250; -x_35 = x_288; -x_36 = x_261; -x_37 = x_304; -x_38 = x_228; -x_39 = x_286; -x_40 = x_306; -x_41 = x_293; -x_42 = x_303; -x_43 = x_296; -x_44 = x_245; -x_45 = x_251; -x_46 = x_230; -x_47 = x_257; -x_48 = x_287; -x_49 = x_232; -x_50 = x_254; -x_51 = x_256; -x_52 = x_321; -x_53 = x_261; -x_54 = x_238; -x_55 = x_246; -x_56 = x_282; -x_57 = x_309; -x_58 = x_295; -x_59 = x_318; -x_60 = x_240; -x_61 = x_234; -x_62 = x_235; -x_63 = x_249; -x_64 = x_239; -x_65 = x_284; -x_66 = x_227; -x_67 = x_268; +x_32 = x_285; +x_33 = x_232; +x_34 = x_256; +x_35 = x_234; +x_36 = x_331; +x_37 = x_287; +x_38 = x_306; +x_39 = x_239; +x_40 = x_257; +x_41 = x_295; +x_42 = x_321; +x_43 = x_254; +x_44 = x_247; +x_45 = x_291; +x_46 = x_301; +x_47 = x_249; +x_48 = x_261; +x_49 = x_246; +x_50 = x_235; +x_51 = x_304; +x_52 = x_261; +x_53 = x_240; +x_54 = x_281; +x_55 = x_296; +x_56 = x_286; +x_57 = x_253; +x_58 = x_297; +x_59 = x_309; +x_60 = x_228; +x_61 = x_311; +x_62 = x_318; +x_63 = x_251; +x_64 = x_227; +x_65 = x_248; +x_66 = x_284; +x_67 = x_230; x_68 = x_334; goto block_225; } diff --git a/stage0/stdlib/Lean/Compiler/IR/Basic.c b/stage0/stdlib/Lean/Compiler/IR/Basic.c index d24471044923..0084c4aac970 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Basic.c +++ b/stage0/stdlib/Lean/Compiler/IR/Basic.c @@ -39,8 +39,8 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_jmp_elim___redArg(lean_object*, lean_o 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*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getJPParams___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isize_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_IRType_isDefiniteRef(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_IRType_isVoid(lean_object*); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__9; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_LocalContext_addParams_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -52,6 +52,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvVarId; LEAN_EXPORT lean_object* l_Lean_IR_IRType_tobject_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Alt_ctorIdx___boxed(lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__26; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__32; static lean_object* l_Lean_IR_instReprIRType_repr___closed__15; LEAN_EXPORT uint8_t l_Lean_IR_instBEqJoinPointId_beq(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprArg_repr___closed__1; @@ -77,6 +78,7 @@ static lean_object* l_Lean_IR_instReprIRType_repr___closed__12; static lean_object* l_Lean_IR_instToStringVarId___lam__0___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_IRType_isScalar___boxed(lean_object*); static lean_object* l_Lean_IR_instReprArg_repr___closed__4; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90____boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_eraseJoinPointDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_Const_get_x3f___at___00Lean_IR_LocalContext_isJP_spec__0___redArg(lean_object*, lean_object*); @@ -95,6 +97,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvExpr; LEAN_EXPORT uint8_t l_Std_DTreeMap_Internal_Impl_contains___at___00Lean_IR_mkIndexSet_spec__0___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instHashableVarId_hash___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_getInfo___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instBEqVarId___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_ret_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__5; @@ -116,6 +119,7 @@ static lean_object* l_Lean_IR_modifyJPs___closed__2; static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__8; static lean_object* l_Lean_IR_instReprParam___closed__0; LEAN_EXPORT uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int64_elim(lean_object*, lean_object*, lean_object*, 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*); @@ -130,6 +134,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_AlphaEqv_ctorIdx___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_IR_instReprIRType_repr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint8_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_var_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int16_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprArg_repr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedFnBody; LEAN_EXPORT lean_object* l_Lean_IR_Alt_modifyBodyM___redArg___lam__0(lean_object*, lean_object*); @@ -153,8 +158,8 @@ static lean_object* l_Lean_IR_FnBody_flatten___closed__0; LEAN_EXPORT uint8_t l_Lean_IR_instBEqIRType_beq___lam__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedParam; static lean_object* l_Lean_IR_instReprIRType_repr___closed__11; -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__0(lean_object*, lean_object*); +static lean_object* l_Lean_IR_instReprIRType_repr___closed__36; 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*); @@ -215,6 +220,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Decl_ctorIdx(lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__10; static lean_object* l_Lean_IR_instReprIRType_repr___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_beq___boxed(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_90_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_union_elim(lean_object*, lean_object*, 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*); @@ -231,6 +237,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instHashableVarId; LEAN_EXPORT lean_object* l_Lean_IR_Decl_fdecl_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqVarId_beq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedDecl; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__34; static lean_object* l_Lean_IR_instInhabitedAlt_default__1___closed__0; LEAN_EXPORT uint8_t l_Lean_IR_LocalContext_isJP(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_VarId_alphaEqv(lean_object*, lean_object*, lean_object*); @@ -245,9 +252,9 @@ 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_List_foldl___at___00List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1_spec__1_spec__1_spec__1(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_85_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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_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 lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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*); @@ -283,7 +290,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___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(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*); @@ -308,6 +314,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_float_elim___redArg(lean_object*, lean 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*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__16; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int16_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_JoinPointId_ctorIdx___boxed(lean_object*); 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; @@ -319,6 +326,7 @@ static lean_object* l_Lean_IR_instReprIRType_repr___closed__27; 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*); +static lean_object* l_Lean_IR_instReprIRType_repr___closed__37; LEAN_EXPORT lean_object* l_Lean_IR_LitVal_ctorIdx(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprJoinPointId; static lean_object* l_Lean_IR_instInhabitedDecl_default___closed__1; @@ -326,6 +334,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedIRType; lean_object* l_Std_DTreeMap_Internal_Impl_minView___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0___closed__1; static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__7; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed___boxed(lean_object*); static lean_object* l_Lean_IR_instReprJoinPointId___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getType___boxed(lean_object*, lean_object*); @@ -333,8 +342,10 @@ static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__2; LEAN_EXPORT uint8_t l_Lean_IR_CtorInfo_isRef(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_args_alphaEqv_spec__0___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instAlphaEqvArrayArg___closed__0; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__33; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Alt_modifyBody(lean_object*, lean_object*); +static lean_object* l_Lean_IR_instReprIRType_repr___closed__40; lean_object* l_Id_instMonad___lam__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_getUnboxOpName___closed__0; static lean_object* l_Lean_IR_instReprCtorInfo___closed__0; @@ -349,6 +360,7 @@ LEAN_EXPORT uint8_t l_Lean_IR_FnBody_alphaEqv(lean_object*, lean_object*, lean_o static lean_object* l_Lean_IR_modifyJPs___closed__7; 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; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__38; LEAN_EXPORT uint8_t l_Lean_IR_Decl_isExtern(lean_object*); static lean_object* l_Lean_IR_instBEqFnBody___closed__0; LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_instBEqIRType_beq_spec__1___redArg(lean_object*, lean_object*, lean_object*); @@ -366,6 +378,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_IRType_boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_addParams___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprIRType; static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__11; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign(uint8_t); LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getJPBody___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_reset_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_FnBody_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); @@ -378,11 +391,13 @@ LEAN_EXPORT lean_object* l_Lean_IR_instReprJoinPointId_repr(lean_object*, lean_o static lean_object* l_Lean_IR_instHashableVarId___closed__0; static lean_object* l_Lean_IR_instReprIRType_repr___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_instHashableJoinPointId; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int32_elim___redArg(lean_object*, lean_object*); static lean_object* l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1___closed__4; LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl_erase___at___00Lean_IR_LocalContext_eraseJoinPointDecl_spec__0___redArg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_ctorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Arg_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_instReprIRType_repr___closed__39; LEAN_EXPORT lean_object* l_Lean_IR_Expr_reuse_elim(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__30; LEAN_EXPORT lean_object* l_Lean_IR_VarId_ctorIdx(lean_object*); @@ -392,6 +407,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Expr_fap_elim(lean_object*, lean_object*, lea static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__13; LEAN_EXPORT lean_object* l_Lean_IR_instToStringJoinPointId; LEAN_EXPORT lean_object* l_Lean_IR_IRType_uint32_elim___redArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int64_elim___redArg(lean_object*, lean_object*); 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*); @@ -405,6 +421,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Decl_resultType___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedLitVal; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_FnBody_beq(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_90____boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_IR_Expr_ctor_elim(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__8; LEAN_EXPORT uint8_t l_Array_isEqvAux___at___00Lean_IR_FnBody_alphaEqv_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -413,6 +430,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_instAlphaEqvArg; LEAN_EXPORT lean_object* l_Lean_IR_Expr_sproj_elim(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instBEqIRType; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__35; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_sset_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_instReprParam_repr___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); @@ -429,6 +447,7 @@ 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; +static lean_object* l_Lean_IR_getUnboxOpName___closed__6; 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*); @@ -436,6 +455,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_type___boxed(lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__24; LEAN_EXPORT lean_object* l_Lean_IR_VarId_ctorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int8_elim___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_box_elim___redArg(lean_object*, lean_object*); 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*); @@ -455,6 +475,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_FnBody_jdecl_elim(lean_object*, lean_object*, 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_IRType_int8_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__3; LEAN_EXPORT lean_object* l_List_foldl___at___00Std_Format_joinSep___at___00Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1_spec__1_spec__1(lean_object*, lean_object*, lean_object*); @@ -491,8 +512,8 @@ static lean_object* l_Lean_IR_instReprVarId_repr___redArg___closed__6; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Index_lt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_ctorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_IRType_isPossibleRef___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85____boxed(lean_object**); static lean_object* l_Lean_IR_instReprCtorInfo_repr___redArg___closed__11; LEAN_EXPORT lean_object* l_Lean_IR_instInhabitedExpr_default; static lean_object* l_Lean_IR_modifyJPs___closed__4; @@ -506,9 +527,11 @@ static lean_object* l_Lean_IR_instReprIRType_repr___closed__29; LEAN_EXPORT lean_object* l_Lean_IR_Decl_getInfo(lean_object*); static lean_object* l_Lean_IR_instReprIRType_repr___closed__22; static lean_object* l_Lean_IR_instReprArg_repr___closed__3; +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isize_elim___redArg(lean_object*, lean_object*); static lean_object* l_Lean_IR_modifyJPs___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_Param_ctorIdx(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprVarId_repr(lean_object*, lean_object*); +static lean_object* l_Lean_IR_getUnboxOpName___closed__8; LEAN_EXPORT lean_object* l_Lean_IR_Expr_uproj_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_IR_addParamsRename_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_EXPORT lean_object* l_Lean_IR_LocalContext_eraseJoinPointDecl___boxed(lean_object*, lean_object*); @@ -544,6 +567,7 @@ static lean_object* l_Lean_IR_instBEqArg___closed__0; lean_object* l_Lean_Name_mkStr1(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Expr_ctorIdx(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instReprParam_repr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int32_elim(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -569,6 +593,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; +static lean_object* l_Lean_IR_instReprIRType_repr___closed__41; LEAN_EXPORT uint64_t l_Lean_IR_instHashableVarId_hash(lean_object*); LEAN_EXPORT lean_object* l_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__0___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -595,6 +620,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_LocalContext_getJPBody(lean_object*, lean_obj LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_LocalContext_addParams_spec__0(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_IR_instAlphaEqvArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instReprCtorInfo_repr___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_IR_getUnboxOpName___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_ret_elim(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -1246,12 +1272,42 @@ lean_object* x_14; x_14 = lean_unsigned_to_nat(12u); return x_14; } -default: +case 13: { lean_object* x_15; x_15 = lean_unsigned_to_nat(13u); return x_15; } +case 14: +{ +lean_object* x_16; +x_16 = lean_unsigned_to_nat(14u); +return x_16; +} +case 15: +{ +lean_object* x_17; +x_17 = lean_unsigned_to_nat(15u); +return x_17; +} +case 16: +{ +lean_object* x_18; +x_18 = lean_unsigned_to_nat(16u); +return x_18; +} +case 17: +{ +lean_object* x_19; +x_19 = lean_unsigned_to_nat(17u); +return x_19; +} +default: +{ +lean_object* x_20; +x_20 = lean_unsigned_to_nat(18u); +return x_20; +} } } } @@ -1539,6 +1595,86 @@ x_5 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_4); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int8_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_IRType_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int8_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_IRType_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int16_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_IRType_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int16_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_IRType_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int32_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_IRType_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int32_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_IRType_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int64_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_IRType_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_int64_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_IRType_ctorElim___redArg(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isize_elim___redArg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_IR_IRType_ctorElim___redArg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_isize_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_IRType_ctorElim___redArg(x_2, x_4); +return x_5; +} +} static lean_object* _init_l_Lean_IR_instInhabitedIRType_default() { _start: { @@ -1555,7 +1691,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(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_IRType_match__on__same__ctor_het___redArg___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(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; @@ -1563,7 +1699,7 @@ x_6 = lean_apply_4(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(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_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(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; @@ -1571,13 +1707,18 @@ x_6 = lean_apply_4(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_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_90_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_17; +lean_object* x_22; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1591,12 +1732,17 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_17 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_3); -return x_17; +x_22 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_3); +return x_22; } case 1: { -lean_object* x_18; +lean_object* x_23; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1610,12 +1756,17 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_18 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_4); -return x_18; +x_23 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_4); +return x_23; } case 2: { -lean_object* x_19; +lean_object* x_24; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1629,12 +1780,17 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_19 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_5); -return x_19; +x_24 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_5); +return x_24; } case 3: { -lean_object* x_20; +lean_object* x_25; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1648,12 +1804,17 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_20 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_6); -return x_20; +x_25 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_6); +return x_25; } case 4: { -lean_object* x_21; +lean_object* x_26; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1667,12 +1828,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_21 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_7); -return x_21; +x_26 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_7); +return x_26; } case 5: { -lean_object* x_22; +lean_object* x_27; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1686,12 +1852,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_22 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_8); -return x_22; +x_27 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_8); +return x_27; } case 6: { -lean_object* x_23; +lean_object* x_28; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1705,12 +1876,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_23 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_9); -return x_23; +x_28 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_9); +return x_28; } case 7: { -lean_object* x_24; +lean_object* x_29; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1724,12 +1900,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_24 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_10); -return x_24; +x_29 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_10); +return x_29; } case 8: { -lean_object* x_25; +lean_object* x_30; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1743,12 +1924,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_25 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_11); -return x_25; +x_30 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_11); +return x_30; } case 9: { -lean_object* x_26; +lean_object* x_31; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1762,12 +1948,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_26 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_12); -return x_26; +x_31 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_12); +return x_31; } case 10: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -1781,21 +1972,26 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_27 = lean_ctor_get(x_1, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_28); +x_32 = lean_ctor_get(x_1, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_33); lean_dec_ref(x_1); -x_29 = lean_alloc_closure((void*)(l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_), 5, 3); -lean_closure_set(x_29, 0, x_13); -lean_closure_set(x_29, 1, x_27); -lean_closure_set(x_29, 2, x_28); -x_30 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_29); -return x_30; +x_34 = lean_alloc_closure((void*)(l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__0_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_), 5, 3); +lean_closure_set(x_34, 0, x_13); +lean_closure_set(x_34, 1, x_32); +lean_closure_set(x_34, 2, x_33); +x_35 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_34); +return x_35; } case 11: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_13); @@ -1809,21 +2005,26 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_31 = lean_ctor_get(x_1, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_32); +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_37); lean_dec_ref(x_1); -x_33 = lean_alloc_closure((void*)(l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_), 5, 3); -lean_closure_set(x_33, 0, x_14); -lean_closure_set(x_33, 1, x_31); -lean_closure_set(x_33, 2, x_32); -x_34 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_33); -return x_34; +x_38 = lean_alloc_closure((void*)(l_Lean_IR_IRType_match__on__same__ctor_het___redArg___lam__1_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_), 5, 3); +lean_closure_set(x_38, 0, x_14); +lean_closure_set(x_38, 1, x_36); +lean_closure_set(x_38, 2, x_37); +x_39 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_38); +return x_39; } case 12: { -lean_object* x_35; +lean_object* x_40; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); @@ -1837,12 +2038,137 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_35 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_15); -return x_35; +x_40 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_15); +return x_40; +} +case 13: +{ +lean_object* x_41; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_41 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_16); +return x_41; +} +case 14: +{ +lean_object* x_42; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_42 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_17); +return x_42; +} +case 15: +{ +lean_object* x_43; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_43 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_18); +return x_43; +} +case 16: +{ +lean_object* x_44; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_44 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_19); +return x_44; +} +case 17: +{ +lean_object* x_45; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_45 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_20); +return x_45; } default: { -lean_object* x_36; +lean_object* x_46; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -1856,21 +2182,50 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_36 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_16); -return x_36; +x_46 = l_Lean_IR_IRType_ctorElim___redArg(x_2, x_21); +return x_46; +} } } } +LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +_start: +{ +lean_object* x_24; +x_24 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(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, x_20, x_21, x_22, x_23); +return x_24; +} } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_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___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90____boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -lean_object* x_19; -x_19 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(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_object* x_22; +x_22 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(x_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); +return x_22; } } -LEAN_EXPORT lean_object* l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85____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_90____boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -1889,11 +2244,16 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; _start: { -lean_object* x_19; -x_19 = l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(x_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); -return x_19; +lean_object* x_24; +x_24 = l_Lean_IR_IRType_match__on__same__ctor_het_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(x_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); +return x_24; } } LEAN_EXPORT uint8_t l_Option_instBEq_beq___at___00Lean_IR_instBEqIRType_beq_spec__0(lean_object* x_1, lean_object* x_2) { @@ -2065,8 +2425,8 @@ x_7 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__1___boxed), x_8 = lean_box(x_5); x_9 = lean_alloc_closure((void*)(l_Lean_IR_instBEqIRType_beq___lam__2___boxed), 3, 1); lean_closure_set(x_9, 0, x_8); -lean_inc_ref_n(x_9, 11); -x_10 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_85_(x_1, x_2, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_7, x_6, x_9, x_9); +lean_inc_ref_n(x_9, 16); +x_10 = l_Lean_IR_IRType_match__on__same__ctor_het___redArg_00___x40_Lean_Compiler_IR_Basic_840659257____hygCtx___hyg_90_(x_1, x_2, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_9, x_7, x_6, x_9, x_9, x_9, x_9, x_9, x_9, x_9); x_11 = lean_apply_2(x_10, lean_box(0), lean_box(0)); x_12 = lean_unbox(x_11); return x_12; @@ -2741,13 +3101,103 @@ return x_2; static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__24() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.IRType.int8", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_instReprIRType_repr___closed__24; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.IRType.int16", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_instReprIRType_repr___closed__26; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.IRType.int32", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_instReprIRType_repr___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_Lean_IR_instReprIRType_repr___closed__30() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.IRType.int64", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_instReprIRType_repr___closed__30; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__32() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.IR.IRType.isize", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_instReprIRType_repr___closed__32; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__34() { +_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_instReprIRType_repr___closed__25() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -2756,7 +3206,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__26() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__36() { _start: { lean_object* x_1; @@ -2764,29 +3214,29 @@ x_1 = lean_mk_string_unchecked("Lean.IR.IRType.struct", 21, 21); return x_1; } } -static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__27() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__37() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_instReprIRType_repr___closed__26; +x_1 = l_Lean_IR_instReprIRType_repr___closed__36; 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_instReprIRType_repr___closed__28() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(1); -x_2 = l_Lean_IR_instReprIRType_repr___closed__27; +x_2 = l_Lean_IR_instReprIRType_repr___closed__37; 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_instReprIRType_repr___closed__29() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__39() { _start: { lean_object* x_1; @@ -2794,22 +3244,22 @@ x_1 = lean_mk_string_unchecked("Lean.IR.IRType.union", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__30() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_IR_instReprIRType_repr___closed__29; +x_1 = l_Lean_IR_instReprIRType_repr___closed__39; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__31() { +static lean_object* _init_l_Lean_IR_instReprIRType_repr___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(1); -x_2 = l_Lean_IR_instReprIRType_repr___closed__30; +x_2 = l_Lean_IR_instReprIRType_repr___closed__40; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -2819,376 +3269,476 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_IR_instReprIRType_repr(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_10; lean_object* x_17; lean_object* x_24; lean_object* x_31; lean_object* x_38; lean_object* x_45; lean_object* x_52; lean_object* x_59; lean_object* x_66; lean_object* x_73; lean_object* x_80; +lean_object* x_3; lean_object* x_10; lean_object* x_17; lean_object* x_24; lean_object* x_31; lean_object* x_38; lean_object* x_45; lean_object* x_52; lean_object* x_59; lean_object* x_66; lean_object* x_73; lean_object* x_80; lean_object* x_87; lean_object* x_94; lean_object* x_101; lean_object* x_108; lean_object* x_115; switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_87; uint8_t x_88; -x_87 = lean_unsigned_to_nat(1024u); -x_88 = lean_nat_dec_le(x_87, x_2); -if (x_88 == 0) +lean_object* x_122; uint8_t x_123; +x_122 = lean_unsigned_to_nat(1024u); +x_123 = lean_nat_dec_le(x_122, x_2); +if (x_123 == 0) { -lean_object* x_89; -x_89 = l_Lean_IR_instReprIRType_repr___closed__24; -x_3 = x_89; +lean_object* x_124; +x_124 = l_Lean_IR_instReprIRType_repr___closed__34; +x_3 = x_124; goto block_9; } else { -lean_object* x_90; -x_90 = l_Lean_IR_instReprIRType_repr___closed__25; -x_3 = x_90; +lean_object* x_125; +x_125 = l_Lean_IR_instReprIRType_repr___closed__35; +x_3 = x_125; goto block_9; } } case 1: { -lean_object* x_91; uint8_t x_92; -x_91 = lean_unsigned_to_nat(1024u); -x_92 = lean_nat_dec_le(x_91, x_2); -if (x_92 == 0) +lean_object* x_126; uint8_t x_127; +x_126 = lean_unsigned_to_nat(1024u); +x_127 = lean_nat_dec_le(x_126, x_2); +if (x_127 == 0) { -lean_object* x_93; -x_93 = l_Lean_IR_instReprIRType_repr___closed__24; -x_10 = x_93; +lean_object* x_128; +x_128 = l_Lean_IR_instReprIRType_repr___closed__34; +x_10 = x_128; goto block_16; } else { -lean_object* x_94; -x_94 = l_Lean_IR_instReprIRType_repr___closed__25; -x_10 = x_94; +lean_object* x_129; +x_129 = l_Lean_IR_instReprIRType_repr___closed__35; +x_10 = x_129; goto block_16; } } case 2: { -lean_object* x_95; uint8_t x_96; -x_95 = lean_unsigned_to_nat(1024u); -x_96 = lean_nat_dec_le(x_95, x_2); -if (x_96 == 0) +lean_object* x_130; uint8_t x_131; +x_130 = lean_unsigned_to_nat(1024u); +x_131 = lean_nat_dec_le(x_130, x_2); +if (x_131 == 0) { -lean_object* x_97; -x_97 = l_Lean_IR_instReprIRType_repr___closed__24; -x_17 = x_97; +lean_object* x_132; +x_132 = l_Lean_IR_instReprIRType_repr___closed__34; +x_17 = x_132; goto block_23; } else { -lean_object* x_98; -x_98 = l_Lean_IR_instReprIRType_repr___closed__25; -x_17 = x_98; +lean_object* x_133; +x_133 = l_Lean_IR_instReprIRType_repr___closed__35; +x_17 = x_133; goto block_23; } } case 3: { -lean_object* x_99; uint8_t x_100; -x_99 = lean_unsigned_to_nat(1024u); -x_100 = lean_nat_dec_le(x_99, x_2); -if (x_100 == 0) +lean_object* x_134; uint8_t x_135; +x_134 = lean_unsigned_to_nat(1024u); +x_135 = lean_nat_dec_le(x_134, x_2); +if (x_135 == 0) { -lean_object* x_101; -x_101 = l_Lean_IR_instReprIRType_repr___closed__24; -x_24 = x_101; +lean_object* x_136; +x_136 = l_Lean_IR_instReprIRType_repr___closed__34; +x_24 = x_136; goto block_30; } else { -lean_object* x_102; -x_102 = l_Lean_IR_instReprIRType_repr___closed__25; -x_24 = x_102; +lean_object* x_137; +x_137 = l_Lean_IR_instReprIRType_repr___closed__35; +x_24 = x_137; goto block_30; } } case 4: { -lean_object* x_103; uint8_t x_104; -x_103 = lean_unsigned_to_nat(1024u); -x_104 = lean_nat_dec_le(x_103, x_2); -if (x_104 == 0) +lean_object* x_138; uint8_t x_139; +x_138 = lean_unsigned_to_nat(1024u); +x_139 = lean_nat_dec_le(x_138, x_2); +if (x_139 == 0) { -lean_object* x_105; -x_105 = l_Lean_IR_instReprIRType_repr___closed__24; -x_31 = x_105; +lean_object* x_140; +x_140 = l_Lean_IR_instReprIRType_repr___closed__34; +x_31 = x_140; goto block_37; } else { -lean_object* x_106; -x_106 = l_Lean_IR_instReprIRType_repr___closed__25; -x_31 = x_106; +lean_object* x_141; +x_141 = l_Lean_IR_instReprIRType_repr___closed__35; +x_31 = x_141; goto block_37; } } case 5: { -lean_object* x_107; uint8_t x_108; -x_107 = lean_unsigned_to_nat(1024u); -x_108 = lean_nat_dec_le(x_107, x_2); -if (x_108 == 0) +lean_object* x_142; uint8_t x_143; +x_142 = lean_unsigned_to_nat(1024u); +x_143 = lean_nat_dec_le(x_142, x_2); +if (x_143 == 0) { -lean_object* x_109; -x_109 = l_Lean_IR_instReprIRType_repr___closed__24; -x_38 = x_109; +lean_object* x_144; +x_144 = l_Lean_IR_instReprIRType_repr___closed__34; +x_38 = x_144; goto block_44; } else { -lean_object* x_110; -x_110 = l_Lean_IR_instReprIRType_repr___closed__25; -x_38 = x_110; +lean_object* x_145; +x_145 = l_Lean_IR_instReprIRType_repr___closed__35; +x_38 = x_145; goto block_44; } } case 6: { -lean_object* x_111; uint8_t x_112; -x_111 = lean_unsigned_to_nat(1024u); -x_112 = lean_nat_dec_le(x_111, x_2); -if (x_112 == 0) +lean_object* x_146; uint8_t x_147; +x_146 = lean_unsigned_to_nat(1024u); +x_147 = lean_nat_dec_le(x_146, x_2); +if (x_147 == 0) { -lean_object* x_113; -x_113 = l_Lean_IR_instReprIRType_repr___closed__24; -x_45 = x_113; +lean_object* x_148; +x_148 = l_Lean_IR_instReprIRType_repr___closed__34; +x_45 = x_148; goto block_51; } else { -lean_object* x_114; -x_114 = l_Lean_IR_instReprIRType_repr___closed__25; -x_45 = x_114; +lean_object* x_149; +x_149 = l_Lean_IR_instReprIRType_repr___closed__35; +x_45 = x_149; goto block_51; } } case 7: { -lean_object* x_115; uint8_t x_116; -x_115 = lean_unsigned_to_nat(1024u); -x_116 = lean_nat_dec_le(x_115, x_2); -if (x_116 == 0) +lean_object* x_150; uint8_t x_151; +x_150 = lean_unsigned_to_nat(1024u); +x_151 = lean_nat_dec_le(x_150, x_2); +if (x_151 == 0) { -lean_object* x_117; -x_117 = l_Lean_IR_instReprIRType_repr___closed__24; -x_52 = x_117; +lean_object* x_152; +x_152 = l_Lean_IR_instReprIRType_repr___closed__34; +x_52 = x_152; goto block_58; } else { -lean_object* x_118; -x_118 = l_Lean_IR_instReprIRType_repr___closed__25; -x_52 = x_118; +lean_object* x_153; +x_153 = l_Lean_IR_instReprIRType_repr___closed__35; +x_52 = x_153; goto block_58; } } case 8: { -lean_object* x_119; uint8_t x_120; -x_119 = lean_unsigned_to_nat(1024u); -x_120 = lean_nat_dec_le(x_119, x_2); -if (x_120 == 0) +lean_object* x_154; uint8_t x_155; +x_154 = lean_unsigned_to_nat(1024u); +x_155 = lean_nat_dec_le(x_154, x_2); +if (x_155 == 0) { -lean_object* x_121; -x_121 = l_Lean_IR_instReprIRType_repr___closed__24; -x_59 = x_121; +lean_object* x_156; +x_156 = l_Lean_IR_instReprIRType_repr___closed__34; +x_59 = x_156; goto block_65; } else { -lean_object* x_122; -x_122 = l_Lean_IR_instReprIRType_repr___closed__25; -x_59 = x_122; +lean_object* x_157; +x_157 = l_Lean_IR_instReprIRType_repr___closed__35; +x_59 = x_157; goto block_65; } } case 9: { -lean_object* x_123; uint8_t x_124; -x_123 = lean_unsigned_to_nat(1024u); -x_124 = lean_nat_dec_le(x_123, x_2); -if (x_124 == 0) +lean_object* x_158; uint8_t x_159; +x_158 = lean_unsigned_to_nat(1024u); +x_159 = lean_nat_dec_le(x_158, x_2); +if (x_159 == 0) { -lean_object* x_125; -x_125 = l_Lean_IR_instReprIRType_repr___closed__24; -x_66 = x_125; +lean_object* x_160; +x_160 = l_Lean_IR_instReprIRType_repr___closed__34; +x_66 = x_160; goto block_72; } else { -lean_object* x_126; -x_126 = l_Lean_IR_instReprIRType_repr___closed__25; -x_66 = x_126; +lean_object* x_161; +x_161 = l_Lean_IR_instReprIRType_repr___closed__35; +x_66 = x_161; 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; -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); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_179; uint8_t x_180; +x_162 = lean_ctor_get(x_1, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_163); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); - x_129 = x_1; + x_164 = x_1; } else { lean_dec_ref(x_1); - x_129 = lean_box(0); + x_164 = 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_179 = lean_unsigned_to_nat(1024u); +x_180 = lean_nat_dec_le(x_179, x_2); +if (x_180 == 0) { -lean_object* x_146; -x_146 = l_Lean_IR_instReprIRType_repr___closed__24; -x_130 = x_146; -goto block_143; +lean_object* x_181; +x_181 = l_Lean_IR_instReprIRType_repr___closed__34; +x_165 = x_181; +goto block_178; } else { -lean_object* x_147; -x_147 = l_Lean_IR_instReprIRType_repr___closed__25; -x_130 = x_147; -goto block_143; -} -block_143: -{ -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); +lean_object* x_182; +x_182 = l_Lean_IR_instReprIRType_repr___closed__35; +x_165 = x_182; +goto block_178; +} +block_178: +{ +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; uint8_t x_175; lean_object* x_176; lean_object* x_177; +x_166 = lean_box(1); +x_167 = l_Lean_IR_instReprIRType_repr___closed__38; +x_168 = lean_unsigned_to_nat(1024u); +x_169 = l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0(x_162, x_168); +if (lean_is_scalar(x_164)) { + x_170 = 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); -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_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_170 = x_164; + lean_ctor_set_tag(x_170, 5); +} +lean_ctor_set(x_170, 0, x_167); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_166); +x_172 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_163); +x_173 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_173, 0, x_171); +lean_ctor_set(x_173, 1, x_172); +x_174 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_174, 0, x_165); +lean_ctor_set(x_174, 1, x_173); +x_175 = 0; +x_176 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set_uint8(x_176, sizeof(void*)*1, x_175); +x_177 = l_Repr_addAppParen(x_176, x_2); +return x_177; } } 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_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_200; uint8_t x_201; +x_183 = lean_ctor_get(x_1, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_184); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); - x_150 = x_1; + x_185 = x_1; } else { lean_dec_ref(x_1); - x_150 = lean_box(0); + x_185 = 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_200 = lean_unsigned_to_nat(1024u); +x_201 = lean_nat_dec_le(x_200, x_2); +if (x_201 == 0) { -lean_object* x_167; -x_167 = l_Lean_IR_instReprIRType_repr___closed__24; -x_151 = x_167; -goto block_164; +lean_object* x_202; +x_202 = l_Lean_IR_instReprIRType_repr___closed__34; +x_186 = x_202; +goto block_199; } else { -lean_object* x_168; -x_168 = l_Lean_IR_instReprIRType_repr___closed__25; -x_151 = x_168; -goto block_164; +lean_object* x_203; +x_203 = l_Lean_IR_instReprIRType_repr___closed__35; +x_186 = x_203; +goto block_199; } -block_164: +block_199: { -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_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; lean_object* x_197; lean_object* x_198; +x_187 = lean_box(1); +x_188 = l_Lean_IR_instReprIRType_repr___closed__41; +x_189 = lean_unsigned_to_nat(1024u); +x_190 = l_Lean_Name_reprPrec(x_183, x_189); +if (lean_is_scalar(x_185)) { + x_191 = lean_alloc_ctor(5, 2, 0); } else { - x_156 = x_150; - lean_ctor_set_tag(x_156, 5); + x_191 = x_185; + lean_ctor_set_tag(x_191, 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; +lean_ctor_set(x_191, 0, x_188); +lean_ctor_set(x_191, 1, x_190); +x_192 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_187); +x_193 = l_Array_Array_repr___at___00Lean_IR_instReprIRType_repr_spec__1(x_184); +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 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_195, 0, x_186); +lean_ctor_set(x_195, 1, x_194); +x_196 = 0; +x_197 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set_uint8(x_197, sizeof(void*)*1, x_196); +x_198 = l_Repr_addAppParen(x_197, x_2); +return x_198; } } 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_204; uint8_t x_205; +x_204 = lean_unsigned_to_nat(1024u); +x_205 = lean_nat_dec_le(x_204, x_2); +if (x_205 == 0) { -lean_object* x_171; -x_171 = l_Lean_IR_instReprIRType_repr___closed__24; -x_73 = x_171; +lean_object* x_206; +x_206 = l_Lean_IR_instReprIRType_repr___closed__34; +x_73 = x_206; goto block_79; } else { -lean_object* x_172; -x_172 = l_Lean_IR_instReprIRType_repr___closed__25; -x_73 = x_172; +lean_object* x_207; +x_207 = l_Lean_IR_instReprIRType_repr___closed__35; +x_73 = x_207; goto block_79; } } -default: +case 13: { -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_208; uint8_t x_209; +x_208 = lean_unsigned_to_nat(1024u); +x_209 = lean_nat_dec_le(x_208, x_2); +if (x_209 == 0) { -lean_object* x_175; -x_175 = l_Lean_IR_instReprIRType_repr___closed__24; -x_80 = x_175; +lean_object* x_210; +x_210 = l_Lean_IR_instReprIRType_repr___closed__34; +x_80 = x_210; goto block_86; } else { -lean_object* x_176; -x_176 = l_Lean_IR_instReprIRType_repr___closed__25; -x_80 = x_176; +lean_object* x_211; +x_211 = l_Lean_IR_instReprIRType_repr___closed__35; +x_80 = x_211; goto block_86; } } +case 14: +{ +lean_object* x_212; uint8_t x_213; +x_212 = lean_unsigned_to_nat(1024u); +x_213 = lean_nat_dec_le(x_212, x_2); +if (x_213 == 0) +{ +lean_object* x_214; +x_214 = l_Lean_IR_instReprIRType_repr___closed__34; +x_87 = x_214; +goto block_93; +} +else +{ +lean_object* x_215; +x_215 = l_Lean_IR_instReprIRType_repr___closed__35; +x_87 = x_215; +goto block_93; +} +} +case 15: +{ +lean_object* x_216; uint8_t x_217; +x_216 = lean_unsigned_to_nat(1024u); +x_217 = lean_nat_dec_le(x_216, x_2); +if (x_217 == 0) +{ +lean_object* x_218; +x_218 = l_Lean_IR_instReprIRType_repr___closed__34; +x_94 = x_218; +goto block_100; +} +else +{ +lean_object* x_219; +x_219 = l_Lean_IR_instReprIRType_repr___closed__35; +x_94 = x_219; +goto block_100; +} +} +case 16: +{ +lean_object* x_220; uint8_t x_221; +x_220 = lean_unsigned_to_nat(1024u); +x_221 = lean_nat_dec_le(x_220, x_2); +if (x_221 == 0) +{ +lean_object* x_222; +x_222 = l_Lean_IR_instReprIRType_repr___closed__34; +x_101 = x_222; +goto block_107; +} +else +{ +lean_object* x_223; +x_223 = l_Lean_IR_instReprIRType_repr___closed__35; +x_101 = x_223; +goto block_107; +} +} +case 17: +{ +lean_object* x_224; uint8_t x_225; +x_224 = lean_unsigned_to_nat(1024u); +x_225 = lean_nat_dec_le(x_224, x_2); +if (x_225 == 0) +{ +lean_object* x_226; +x_226 = l_Lean_IR_instReprIRType_repr___closed__34; +x_108 = x_226; +goto block_114; +} +else +{ +lean_object* x_227; +x_227 = l_Lean_IR_instReprIRType_repr___closed__35; +x_108 = x_227; +goto block_114; +} +} +default: +{ +lean_object* x_228; uint8_t x_229; +x_228 = lean_unsigned_to_nat(1024u); +x_229 = lean_nat_dec_le(x_228, x_2); +if (x_229 == 0) +{ +lean_object* x_230; +x_230 = l_Lean_IR_instReprIRType_repr___closed__34; +x_115 = x_230; +goto block_121; +} +else +{ +lean_object* x_231; +x_231 = l_Lean_IR_instReprIRType_repr___closed__35; +x_115 = x_231; +goto block_121; +} +} } block_9: { @@ -3358,6 +3908,76 @@ lean_ctor_set_uint8(x_84, sizeof(void*)*1, x_83); x_85 = l_Repr_addAppParen(x_84, x_2); return x_85; } +block_93: +{ +lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; +x_88 = l_Lean_IR_instReprIRType_repr___closed__25; +x_89 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = 0; +x_91 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set_uint8(x_91, sizeof(void*)*1, x_90); +x_92 = l_Repr_addAppParen(x_91, x_2); +return x_92; +} +block_100: +{ +lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +x_95 = l_Lean_IR_instReprIRType_repr___closed__27; +x_96 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = 0; +x_98 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set_uint8(x_98, sizeof(void*)*1, x_97); +x_99 = l_Repr_addAppParen(x_98, x_2); +return x_99; +} +block_107: +{ +lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; +x_102 = l_Lean_IR_instReprIRType_repr___closed__29; +x_103 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +x_104 = 0; +x_105 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set_uint8(x_105, sizeof(void*)*1, x_104); +x_106 = l_Repr_addAppParen(x_105, x_2); +return x_106; +} +block_114: +{ +lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; +x_109 = l_Lean_IR_instReprIRType_repr___closed__31; +x_110 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +x_111 = 0; +x_112 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set_uint8(x_112, sizeof(void*)*1, x_111); +x_113 = l_Repr_addAppParen(x_112, x_2); +return x_113; +} +block_121: +{ +lean_object* x_116; lean_object* x_117; uint8_t x_118; lean_object* x_119; lean_object* x_120; +x_116 = l_Lean_IR_instReprIRType_repr___closed__33; +x_117 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +x_118 = 0; +x_119 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set_uint8(x_119, sizeof(void*)*1, x_118); +x_120 = l_Repr_addAppParen(x_119, x_2); +return x_120; +} } } LEAN_EXPORT lean_object* l_Option_repr___at___00Lean_IR_instReprIRType_repr_spec__0___boxed(lean_object* x_1, lean_object* x_2) { @@ -3686,6 +4306,32 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_box(5); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_box(18); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +x_3 = l_Lean_IR_IRType_ptrSizedTypeForSign(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_IR_Arg_ctorIdx(lean_object* x_1) { _start: { @@ -3932,14 +4578,14 @@ x_21 = lean_nat_dec_le(x_20, x_2); if (x_21 == 0) { lean_object* x_22; -x_22 = l_Lean_IR_instReprIRType_repr___closed__24; +x_22 = l_Lean_IR_instReprIRType_repr___closed__34; x_11 = x_22; goto block_19; } else { lean_object* x_23; -x_23 = l_Lean_IR_instReprIRType_repr___closed__25; +x_23 = l_Lean_IR_instReprIRType_repr___closed__35; x_11 = x_23; goto block_19; } @@ -3970,14 +4616,14 @@ x_25 = lean_nat_dec_le(x_24, x_2); if (x_25 == 0) { lean_object* x_26; -x_26 = l_Lean_IR_instReprIRType_repr___closed__24; +x_26 = l_Lean_IR_instReprIRType_repr___closed__34; x_3 = x_26; goto block_9; } else { lean_object* x_27; -x_27 = l_Lean_IR_instReprIRType_repr___closed__25; +x_27 = l_Lean_IR_instReprIRType_repr___closed__35; x_3 = x_27; goto block_9; } @@ -4934,89 +5580,102 @@ x_10 = lean_box(x_8); x_11 = lean_apply_4(x_2, x_6, x_7, x_10, x_9); return x_11; } -case 5: +case 4: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +x_14 = lean_ctor_get(x_1, 1); lean_inc(x_14); lean_dec_ref(x_1); -x_15 = lean_apply_3(x_2, x_12, x_13, x_14); -return x_15; +x_15 = lean_box(x_13); +x_16 = lean_apply_3(x_2, x_12, x_15, x_14); +return x_16; +} +case 5: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +lean_dec_ref(x_1); +x_20 = lean_apply_3(x_2, x_17, x_18, x_19); +return x_20; } case 6: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -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_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_1, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_22); lean_dec_ref(x_1); -x_18 = lean_apply_2(x_2, x_16, x_17); -return x_18; +x_23 = lean_apply_2(x_2, x_21, x_22); +return x_23; } case 7: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_1, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_25); lean_dec_ref(x_1); -x_21 = lean_apply_2(x_2, x_19, x_20); -return x_21; +x_26 = lean_apply_2(x_2, x_24, x_25); +return x_26; } case 8: { -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_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_28); lean_dec_ref(x_1); -x_24 = lean_apply_2(x_2, x_22, x_23); -return x_24; +x_29 = lean_apply_2(x_2, x_27, x_28); +return x_29; } 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_30; lean_object* x_31; +x_30 = lean_ctor_get(x_1, 0); +lean_inc(x_30); lean_dec_ref(x_1); -x_26 = lean_apply_1(x_2, x_25); -return x_26; +x_31 = lean_apply_1(x_2, x_30); +return x_31; } 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_32; lean_object* x_33; +x_32 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_32); lean_dec_ref(x_1); -x_28 = lean_apply_1(x_2, x_27); -return x_28; +x_33 = lean_apply_1(x_2, x_32); +return x_33; } 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_34; lean_object* x_35; +x_34 = lean_ctor_get(x_1, 0); +lean_inc(x_34); lean_dec_ref(x_1); -x_30 = lean_apply_1(x_2, x_29); -return x_30; +x_35 = lean_apply_1(x_2, x_34); +return x_35; } 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_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_33 = lean_apply_2(x_2, x_31, x_32); -return x_33; +x_38 = lean_apply_2(x_2, x_36, x_37); +return x_38; } } } @@ -5788,117 +6447,115 @@ 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; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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); -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_25 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_26 = lean_ctor_get(x_1, 3); +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_box(x_25); +x_28 = lean_apply_5(x_2, x_22, x_23, x_24, x_27, x_26); +return x_28; } -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_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_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, 1); +x_33 = lean_ctor_get(x_1, 4); 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, 5); +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_6(x_2, 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: { @@ -5906,18 +6563,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; } } } @@ -6489,142 +7148,144 @@ return x_1; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; x_28 = lean_ctor_get(x_1, 0); x_29 = lean_ctor_get(x_1, 1); x_30 = lean_ctor_get(x_1, 2); +x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); lean_dec(x_1); -x_31 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_31, 0, x_28); -lean_ctor_set(x_31, 1, x_29); -lean_ctor_set(x_31, 2, x_30); -lean_ctor_set(x_31, 3, x_2); -return x_31; +x_32 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_32, 0, x_28); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set(x_32, 3, x_2); +lean_ctor_set_uint8(x_32, sizeof(void*)*4, x_31); +return x_32; } } case 5: { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_1); -if (x_32 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_1); +if (x_33 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_1, 5); -lean_dec(x_33); +lean_object* x_34; +x_34 = lean_ctor_get(x_1, 5); +lean_dec(x_34); lean_ctor_set(x_1, 5, x_2); return x_1; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_34 = lean_ctor_get(x_1, 0); -x_35 = lean_ctor_get(x_1, 1); -x_36 = lean_ctor_get(x_1, 2); -x_37 = lean_ctor_get(x_1, 3); -x_38 = lean_ctor_get(x_1, 4); +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_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); +x_39 = lean_ctor_get(x_1, 4); +lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); lean_dec(x_1); -x_39 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_35); -lean_ctor_set(x_39, 2, x_36); -lean_ctor_set(x_39, 3, x_37); -lean_ctor_set(x_39, 4, x_38); -lean_ctor_set(x_39, 5, x_2); -return x_39; +x_40 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_40, 0, x_35); +lean_ctor_set(x_40, 1, x_36); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_38); +lean_ctor_set(x_40, 4, x_39); +lean_ctor_set(x_40, 5, x_2); +return x_40; } } case 6: { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_1); -if (x_40 == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_1); +if (x_41 == 0) { -lean_object* x_41; -x_41 = lean_ctor_get(x_1, 2); -lean_dec(x_41); +lean_object* x_42; +x_42 = lean_ctor_get(x_1, 2); +lean_dec(x_42); 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_object* x_43; lean_object* x_44; uint8_t x_45; uint8_t x_46; lean_object* x_47; +x_43 = lean_ctor_get(x_1, 0); +x_44 = lean_ctor_get(x_1, 1); +x_45 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); 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_47 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_47, 0, x_43); +lean_ctor_set(x_47, 1, x_44); +lean_ctor_set(x_47, 2, x_2); +lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*3 + 1, x_46); +return x_47; } } case 7: { -uint8_t x_47; -x_47 = !lean_is_exclusive(x_1); -if (x_47 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_1); +if (x_48 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_1, 2); -lean_dec(x_48); +lean_object* x_49; +x_49 = lean_ctor_get(x_1, 2); +lean_dec(x_49); 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_object* x_50; lean_object* x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; +x_50 = lean_ctor_get(x_1, 0); +x_51 = lean_ctor_get(x_1, 1); +x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_53 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); 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_54 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_54, 0, x_50); +lean_ctor_set(x_54, 1, x_51); +lean_ctor_set(x_54, 2, x_2); +lean_ctor_set_uint8(x_54, sizeof(void*)*3, x_52); +lean_ctor_set_uint8(x_54, sizeof(void*)*3 + 1, x_53); +return x_54; } } case 8: { -uint8_t x_54; -x_54 = !lean_is_exclusive(x_1); -if (x_54 == 0) +uint8_t x_55; +x_55 = !lean_is_exclusive(x_1); +if (x_55 == 0) { -lean_object* x_55; -x_55 = lean_ctor_get(x_1, 1); -lean_dec(x_55); +lean_object* x_56; +x_56 = lean_ctor_get(x_1, 1); +lean_dec(x_56); 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_57; lean_object* x_58; +x_57 = lean_ctor_get(x_1, 0); +lean_inc(x_57); 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_58 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_2); +return x_58; } } default: @@ -7830,7 +8491,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(390u); 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); @@ -12584,11 +13245,11 @@ 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_11 = x_25; +x_12 = x_26; +x_13 = x_27; +x_14 = x_28; +goto block_17; } else { @@ -12710,11 +13371,11 @@ lean_inc(x_47); x_48 = lean_ctor_get(x_3, 1); lean_inc(x_48); lean_dec_ref(x_3); -x_4 = x_45; -x_5 = x_46; -x_6 = x_47; -x_7 = x_48; -goto block_10; +x_11 = x_45; +x_12 = x_46; +x_13 = x_47; +x_14 = x_48; +goto block_17; } else { @@ -12729,354 +13390,398 @@ 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; +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; uint8_t x_59; 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); +x_51 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_52 = lean_ctor_get(x_2, 1); lean_inc(x_52); -x_53 = lean_ctor_get(x_3, 1); +lean_dec_ref(x_2); +x_53 = lean_ctor_get(x_3, 0); lean_inc(x_53); +x_54 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_55 = lean_ctor_get(x_3, 1); +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; -goto block_10; +x_59 = lean_nat_dec_eq(x_50, x_53); +lean_dec(x_53); +lean_dec(x_50); +if (x_59 == 0) +{ +x_56 = x_59; +goto block_58; +} +else +{ +if (x_51 == 0) +{ +if (x_54 == 0) +{ +x_56 = x_59; +goto block_58; +} +else +{ +lean_dec(x_55); +lean_dec(x_52); +return x_51; +} +} +else +{ +x_56 = x_54; +goto block_58; +} +} +block_58: +{ +if (x_56 == 0) +{ +lean_dec(x_55); +lean_dec(x_52); +return x_56; +} +else +{ +uint8_t x_57; +x_57 = l_Lean_IR_VarId_alphaEqv(x_1, x_52, x_55); +lean_dec(x_55); +lean_dec(x_52); +return x_57; +} +} } else { -uint8_t x_54; +uint8_t x_60; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_54 = 0; -return x_54; +x_60 = 0; +return x_60; } } 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_inc(x_57); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; uint8_t x_70; +x_61 = lean_ctor_get(x_2, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_2, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_2, 2); +lean_inc(x_63); lean_dec_ref(x_2); -x_58 = lean_ctor_get(x_3, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_3, 1); -lean_inc(x_59); -x_60 = lean_ctor_get(x_3, 2); -lean_inc(x_60); +x_64 = lean_ctor_get(x_3, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_3, 1); +lean_inc(x_65); +x_66 = lean_ctor_get(x_3, 2); +lean_inc(x_66); 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_70 = lean_nat_dec_eq(x_61, x_64); +lean_dec(x_64); +lean_dec(x_61); +if (x_70 == 0) { -lean_dec(x_59); -lean_dec(x_56); -x_61 = x_64; -goto block_63; +lean_dec(x_65); +lean_dec(x_62); +x_67 = x_70; +goto block_69; } 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_71; +x_71 = lean_nat_dec_eq(x_62, x_65); +lean_dec(x_65); +lean_dec(x_62); +x_67 = x_71; +goto block_69; } -block_63: +block_69: { -if (x_61 == 0) +if (x_67 == 0) { -lean_dec(x_60); -lean_dec(x_57); -return x_61; +lean_dec(x_66); +lean_dec(x_63); +return x_67; } 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_68; +x_68 = l_Lean_IR_VarId_alphaEqv(x_1, x_63, x_66); +lean_dec(x_66); +lean_dec(x_63); +return x_68; } } } else { -uint8_t x_66; +uint8_t x_72; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_66 = 0; -return x_66; +x_72 = 0; +return x_72; } } 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_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_2, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_74); lean_dec_ref(x_2); -x_69 = lean_ctor_get(x_3, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_70); +x_75 = lean_ctor_get(x_3, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_76); lean_dec_ref(x_3); -x_11 = x_67; -x_12 = x_68; -x_13 = x_69; -x_14 = x_70; -goto block_17; +x_4 = x_73; +x_5 = x_74; +x_6 = x_75; +x_7 = x_76; +goto block_10; } else { -uint8_t x_71; +uint8_t x_77; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_71 = 0; -return x_71; +x_77 = 0; +return x_77; } } 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_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_2, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_79); lean_dec_ref(x_2); -x_74 = lean_ctor_get(x_3, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_75); +x_80 = lean_ctor_get(x_3, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_81); lean_dec_ref(x_3); -x_11 = x_72; -x_12 = x_73; -x_13 = x_74; -x_14 = x_75; -goto block_17; +x_4 = x_78; +x_5 = x_79; +x_6 = x_80; +x_7 = x_81; +goto block_10; } else { -uint8_t x_76; +uint8_t x_82; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_76 = 0; -return x_76; +x_82 = 0; +return x_82; } } 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_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_83 = lean_ctor_get(x_2, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_84); lean_dec_ref(x_2); -x_79 = lean_ctor_get(x_3, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_80); +x_85 = lean_ctor_get(x_3, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_86); lean_dec_ref(x_3); -x_81 = l_Lean_IR_VarId_alphaEqv(x_1, x_77, x_79); -lean_dec(x_79); -lean_dec(x_77); -if (x_81 == 0) +x_87 = l_Lean_IR_VarId_alphaEqv(x_1, x_83, x_85); +lean_dec(x_85); +lean_dec(x_83); +if (x_87 == 0) { -lean_dec_ref(x_80); -lean_dec_ref(x_78); -return x_81; +lean_dec_ref(x_86); +lean_dec_ref(x_84); +return x_87; } else { -uint8_t x_82; -x_82 = l_Lean_IR_args_alphaEqv(x_1, x_78, x_80); -lean_dec_ref(x_80); -lean_dec_ref(x_78); -return x_82; +uint8_t x_88; +x_88 = l_Lean_IR_args_alphaEqv(x_1, x_84, x_86); +lean_dec_ref(x_86); +lean_dec_ref(x_84); +return x_88; } } else { -uint8_t x_83; +uint8_t x_89; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_83 = 0; -return x_83; +x_89 = 0; +return x_89; } } 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_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_90 = lean_ctor_get(x_2, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_2, 1); +lean_inc(x_91); lean_dec_ref(x_2); -x_86 = lean_ctor_get(x_3, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_3, 1); -lean_inc(x_87); +x_92 = lean_ctor_get(x_3, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_3, 1); +lean_inc(x_93); lean_dec_ref(x_3); -x_88 = l_Lean_IR_instBEqIRType_beq(x_84, x_86); -if (x_88 == 0) +x_94 = l_Lean_IR_instBEqIRType_beq(x_90, x_92); +if (x_94 == 0) { -lean_dec(x_87); -lean_dec(x_85); -return x_88; +lean_dec(x_93); +lean_dec(x_91); +return x_94; } else { -uint8_t x_89; -x_89 = l_Lean_IR_VarId_alphaEqv(x_1, x_85, x_87); -lean_dec(x_87); -lean_dec(x_85); -return x_89; +uint8_t x_95; +x_95 = l_Lean_IR_VarId_alphaEqv(x_1, x_91, x_93); +lean_dec(x_93); +lean_dec(x_91); +return x_95; } } else { -uint8_t x_90; +uint8_t x_96; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_90 = 0; -return x_90; +x_96 = 0; +return x_96; } } 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_97; lean_object* x_98; uint8_t x_99; +x_97 = lean_ctor_get(x_2, 0); +lean_inc(x_97); lean_dec_ref(x_2); -x_92 = lean_ctor_get(x_3, 0); -lean_inc(x_92); +x_98 = lean_ctor_get(x_3, 0); +lean_inc(x_98); 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_99 = l_Lean_IR_VarId_alphaEqv(x_1, x_97, x_98); +lean_dec(x_98); +lean_dec(x_97); +return x_99; } else { -uint8_t x_94; +uint8_t x_100; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_94 = 0; -return x_94; +x_100 = 0; +return x_100; } } 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_101; lean_object* x_102; uint8_t x_103; +x_101 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_101); lean_dec_ref(x_2); -x_96 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_96); +x_102 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_102); 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_103 = l_Lean_IR_instBEqLitVal_beq(x_101, x_102); +lean_dec_ref(x_102); +lean_dec_ref(x_101); +return x_103; } else { -uint8_t x_98; +uint8_t x_104; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_98 = 0; -return x_98; +x_104 = 0; +return x_104; } } 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_105; lean_object* x_106; uint8_t x_107; +x_105 = lean_ctor_get(x_2, 0); +lean_inc(x_105); lean_dec_ref(x_2); -x_100 = lean_ctor_get(x_3, 0); -lean_inc(x_100); -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_106 = lean_ctor_get(x_3, 0); +lean_inc(x_106); +lean_dec_ref(x_3); +x_107 = l_Lean_IR_VarId_alphaEqv(x_1, x_105, x_106); +lean_dec(x_106); +lean_dec(x_105); +return x_107; } else { -uint8_t x_102; +uint8_t x_108; lean_dec_ref(x_3); lean_dec_ref(x_2); -x_102 = 0; -return x_102; +x_108 = 0; +return x_108; } } } block_10: { uint8_t x_8; -x_8 = lean_nat_dec_eq(x_4, x_6); +x_8 = lean_name_eq(x_4, x_6); lean_dec(x_6); lean_dec(x_4); if (x_8 == 0) { -lean_dec(x_7); -lean_dec(x_5); +lean_dec_ref(x_7); +lean_dec_ref(x_5); return x_8; } else { uint8_t x_9; -x_9 = l_Lean_IR_VarId_alphaEqv(x_1, x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); +x_9 = l_Lean_IR_args_alphaEqv(x_1, x_5, x_7); +lean_dec_ref(x_7); +lean_dec_ref(x_5); return x_9; } } block_17: { uint8_t x_15; -x_15 = lean_name_eq(x_11, x_13); +x_15 = lean_nat_dec_eq(x_11, x_13); lean_dec(x_13); lean_dec(x_11); if (x_15 == 0) { -lean_dec_ref(x_14); -lean_dec_ref(x_12); +lean_dec(x_14); +lean_dec(x_12); return x_15; } else { uint8_t x_16; -x_16 = l_Lean_IR_args_alphaEqv(x_1, x_12, x_14); -lean_dec_ref(x_14); -lean_dec_ref(x_12); +x_16 = l_Lean_IR_VarId_alphaEqv(x_1, x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); return x_16; } } @@ -13418,7 +14123,7 @@ return x_7; 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; lean_object* x_6; uint8_t x_7; lean_object* x_8; uint8_t x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* 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; +lean_object* x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t 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; switch (lean_obj_tag(x_2)) { case 0: { @@ -13721,196 +14426,227 @@ case 4: { 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; +lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; uint8_t x_102; uint8_t x_107; x_92 = lean_ctor_get(x_2, 0); lean_inc(x_92); x_93 = lean_ctor_get(x_2, 1); lean_inc(x_93); x_94 = lean_ctor_get(x_2, 2); lean_inc(x_94); -x_95 = lean_ctor_get(x_2, 3); -lean_inc(x_95); -lean_dec_ref(x_2); -x_96 = lean_ctor_get(x_3, 0); +x_95 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); +x_96 = lean_ctor_get(x_2, 3); lean_inc(x_96); -x_97 = lean_ctor_get(x_3, 1); +lean_dec_ref(x_2); +x_97 = lean_ctor_get(x_3, 0); lean_inc(x_97); -x_98 = lean_ctor_get(x_3, 2); +x_98 = lean_ctor_get(x_3, 1); lean_inc(x_98); -x_99 = lean_ctor_get(x_3, 3); +x_99 = lean_ctor_get(x_3, 2); lean_inc(x_99); +x_100 = lean_ctor_get_uint8(x_3, sizeof(void*)*4); +x_101 = lean_ctor_get(x_3, 3); +lean_inc(x_101); lean_dec_ref(x_3); -x_104 = l_Lean_IR_VarId_alphaEqv(x_1, x_92, x_96); -lean_dec(x_96); +x_107 = l_Lean_IR_VarId_alphaEqv(x_1, x_92, x_97); +lean_dec(x_97); lean_dec(x_92); -if (x_104 == 0) +if (x_107 == 0) { -lean_dec(x_97); +lean_dec(x_98); lean_dec(x_93); -x_100 = x_104; -goto block_103; +x_102 = x_107; +goto block_106; } else { -uint8_t x_105; -x_105 = lean_nat_dec_eq(x_93, x_97); -lean_dec(x_97); +uint8_t x_108; +x_108 = lean_nat_dec_eq(x_93, x_98); +lean_dec(x_98); lean_dec(x_93); -x_100 = x_105; -goto block_103; +x_102 = x_108; +goto block_106; } -block_103: +block_106: { -if (x_100 == 0) +if (x_102 == 0) { +lean_dec(x_101); lean_dec(x_99); -lean_dec(x_98); -lean_dec(x_95); +lean_dec(x_96); lean_dec(x_94); lean_dec(x_1); -return x_100; +return x_102; } else { -uint8_t x_101; -x_101 = l_Lean_IR_VarId_alphaEqv(x_1, x_94, x_98); -lean_dec(x_98); +uint8_t x_103; +x_103 = l_Lean_IR_VarId_alphaEqv(x_1, x_94, x_99); +lean_dec(x_99); lean_dec(x_94); -if (x_101 == 0) +if (x_103 == 0) { -lean_dec(x_99); -lean_dec(x_95); +lean_dec(x_101); +lean_dec(x_96); lean_dec(x_1); -return x_101; +return x_103; +} +else +{ +if (x_95 == 0) +{ +if (x_100 == 0) +{ +x_2 = x_96; +x_3 = x_101; +goto _start; +} +else +{ +lean_dec(x_101); +lean_dec(x_96); +lean_dec(x_1); +return x_95; +} +} +else +{ +if (x_100 == 0) +{ +lean_dec(x_101); +lean_dec(x_96); +lean_dec(x_1); +return x_100; } else { -x_2 = x_95; -x_3 = x_99; +x_2 = x_96; +x_3 = x_101; goto _start; } } } } +} +} else { -uint8_t x_106; +uint8_t x_109; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_106 = 0; -return x_106; +x_109 = 0; +return x_109; } } case 5: { 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); -lean_inc(x_108); -x_109 = lean_ctor_get(x_2, 2); -lean_inc(x_109); -x_110 = lean_ctor_get(x_2, 3); +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; uint8_t x_122; uint8_t x_123; uint8_t x_128; +x_110 = lean_ctor_get(x_2, 0); lean_inc(x_110); -x_111 = lean_ctor_get(x_2, 4); +x_111 = lean_ctor_get(x_2, 1); lean_inc(x_111); -x_112 = lean_ctor_get(x_2, 5); +x_112 = lean_ctor_get(x_2, 2); lean_inc(x_112); -lean_dec_ref(x_2); -x_113 = lean_ctor_get(x_3, 0); +x_113 = lean_ctor_get(x_2, 3); lean_inc(x_113); -x_114 = lean_ctor_get(x_3, 1); +x_114 = lean_ctor_get(x_2, 4); lean_inc(x_114); -x_115 = lean_ctor_get(x_3, 2); +x_115 = lean_ctor_get(x_2, 5); lean_inc(x_115); -x_116 = lean_ctor_get(x_3, 3); +lean_dec_ref(x_2); +x_116 = lean_ctor_get(x_3, 0); lean_inc(x_116); -x_117 = lean_ctor_get(x_3, 4); +x_117 = lean_ctor_get(x_3, 1); lean_inc(x_117); -x_118 = lean_ctor_get(x_3, 5); +x_118 = lean_ctor_get(x_3, 2); lean_inc(x_118); +x_119 = lean_ctor_get(x_3, 3); +lean_inc(x_119); +x_120 = lean_ctor_get(x_3, 4); +lean_inc(x_120); +x_121 = lean_ctor_get(x_3, 5); +lean_inc(x_121); lean_dec_ref(x_3); -x_119 = lean_nat_dec_eq(x_109, x_115); -lean_dec(x_115); -lean_dec(x_109); -x_125 = l_Lean_IR_VarId_alphaEqv(x_1, x_107, x_113); -lean_dec(x_113); -lean_dec(x_107); -if (x_125 == 0) +x_122 = lean_nat_dec_eq(x_112, x_118); +lean_dec(x_118); +lean_dec(x_112); +x_128 = l_Lean_IR_VarId_alphaEqv(x_1, x_110, x_116); +lean_dec(x_116); +lean_dec(x_110); +if (x_128 == 0) { -lean_dec(x_114); -lean_dec(x_108); -x_120 = x_125; -goto block_124; +lean_dec(x_117); +lean_dec(x_111); +x_123 = x_128; +goto block_127; } else { -uint8_t x_126; -x_126 = lean_nat_dec_eq(x_108, x_114); -lean_dec(x_114); -lean_dec(x_108); -x_120 = x_126; -goto block_124; +uint8_t x_129; +x_129 = lean_nat_dec_eq(x_111, x_117); +lean_dec(x_117); +lean_dec(x_111); +x_123 = x_129; +goto block_127; } -block_124: +block_127: { -if (x_120 == 0) +if (x_123 == 0) { -lean_dec(x_118); -lean_dec(x_117); -lean_dec(x_116); -lean_dec(x_112); -lean_dec(x_111); -lean_dec(x_110); +lean_dec(x_121); +lean_dec(x_120); +lean_dec(x_119); +lean_dec(x_115); +lean_dec(x_114); +lean_dec(x_113); lean_dec(x_1); -return x_120; +return x_123; } else { -if (x_119 == 0) +if (x_122 == 0) { -lean_dec(x_118); -lean_dec(x_117); -lean_dec(x_116); -lean_dec(x_112); -lean_dec(x_111); -lean_dec(x_110); +lean_dec(x_121); +lean_dec(x_120); +lean_dec(x_119); +lean_dec(x_115); +lean_dec(x_114); +lean_dec(x_113); lean_dec(x_1); -return x_119; +return x_122; } else { -uint8_t x_121; -x_121 = l_Lean_IR_VarId_alphaEqv(x_1, x_110, x_116); -lean_dec(x_116); -lean_dec(x_110); -if (x_121 == 0) +uint8_t x_124; +x_124 = l_Lean_IR_VarId_alphaEqv(x_1, x_113, x_119); +lean_dec(x_119); +lean_dec(x_113); +if (x_124 == 0) { -lean_dec(x_118); -lean_dec(x_117); -lean_dec(x_112); -lean_dec(x_111); +lean_dec(x_121); +lean_dec(x_120); +lean_dec(x_115); +lean_dec(x_114); lean_dec(x_1); -return x_121; +return x_124; } else { -uint8_t x_122; -x_122 = l_Lean_IR_instBEqIRType_beq(x_111, x_117); -if (x_122 == 0) +uint8_t x_125; +x_125 = l_Lean_IR_instBEqIRType_beq(x_114, x_120); +if (x_125 == 0) { -lean_dec(x_118); -lean_dec(x_112); +lean_dec(x_121); +lean_dec(x_115); lean_dec(x_1); -return x_122; +return x_125; } else { -x_2 = x_112; -x_3 = x_118; +x_2 = x_115; +x_3 = x_121; goto _start; } } @@ -13920,301 +14656,301 @@ goto _start; } else { -uint8_t x_127; +uint8_t x_130; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_127 = 0; -return x_127; +x_130 = 0; +return x_130; } } case 6: { if (lean_obj_tag(x_3) == 6) { -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; uint8_t x_135; uint8_t x_136; lean_object* x_137; -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_uint8(x_2, sizeof(void*)*3); -x_131 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); -x_132 = lean_ctor_get(x_2, 2); +lean_object* x_131; lean_object* x_132; uint8_t x_133; uint8_t x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; uint8_t x_139; lean_object* x_140; +x_131 = lean_ctor_get(x_2, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_2, 1); lean_inc(x_132); +x_133 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_134 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_135 = lean_ctor_get(x_2, 2); +lean_inc(x_135); lean_dec_ref(x_2); -x_133 = lean_ctor_get(x_3, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_3, 1); -lean_inc(x_134); -x_135 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_136 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); -x_137 = lean_ctor_get(x_3, 2); +x_136 = lean_ctor_get(x_3, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_3, 1); lean_inc(x_137); +x_138 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_139 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); +x_140 = lean_ctor_get(x_3, 2); +lean_inc(x_140); lean_dec_ref(x_3); x_21 = x_1; -x_22 = x_128; -x_23 = x_129; -x_24 = x_130; -x_25 = x_131; -x_26 = x_132; -x_27 = x_133; -x_28 = x_134; -x_29 = x_135; -x_30 = x_136; -x_31 = x_137; +x_22 = x_131; +x_23 = x_132; +x_24 = x_133; +x_25 = x_134; +x_26 = x_135; +x_27 = x_136; +x_28 = x_137; +x_29 = x_138; +x_30 = x_139; +x_31 = x_140; goto block_34; } else { -uint8_t x_138; +uint8_t x_141; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_138 = 0; -return x_138; +x_141 = 0; +return x_141; } } case 7: { if (lean_obj_tag(x_3) == 7) { -lean_object* x_139; lean_object* x_140; uint8_t x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; uint8_t x_147; lean_object* x_148; -x_139 = lean_ctor_get(x_2, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_2, 1); -lean_inc(x_140); -x_141 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); -x_142 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); -x_143 = lean_ctor_get(x_2, 2); +lean_object* x_142; lean_object* x_143; uint8_t x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; uint8_t x_150; lean_object* x_151; +x_142 = lean_ctor_get(x_2, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_2, 1); lean_inc(x_143); +x_144 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_145 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_146 = lean_ctor_get(x_2, 2); +lean_inc(x_146); lean_dec_ref(x_2); -x_144 = lean_ctor_get(x_3, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_3, 1); -lean_inc(x_145); -x_146 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_147 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); -x_148 = lean_ctor_get(x_3, 2); +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_uint8(x_3, sizeof(void*)*3); +x_150 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); +x_151 = lean_ctor_get(x_3, 2); +lean_inc(x_151); lean_dec_ref(x_3); x_21 = x_1; -x_22 = x_139; -x_23 = x_140; -x_24 = x_141; -x_25 = x_142; -x_26 = x_143; -x_27 = x_144; -x_28 = x_145; -x_29 = x_146; -x_30 = x_147; -x_31 = x_148; +x_22 = x_142; +x_23 = x_143; +x_24 = x_144; +x_25 = x_145; +x_26 = x_146; +x_27 = x_147; +x_28 = x_148; +x_29 = x_149; +x_30 = x_150; +x_31 = x_151; goto block_34; } else { -uint8_t x_149; +uint8_t x_152; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_149 = 0; -return x_149; +x_152 = 0; +return x_152; } } case 8: { if (lean_obj_tag(x_3) == 8) { -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_2, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_2, 1); -lean_inc(x_151); -lean_dec_ref(x_2); -x_152 = lean_ctor_get(x_3, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_3, 1); +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_153 = lean_ctor_get(x_2, 0); lean_inc(x_153); +x_154 = lean_ctor_get(x_2, 1); +lean_inc(x_154); +lean_dec_ref(x_2); +x_155 = lean_ctor_get(x_3, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_3, 1); +lean_inc(x_156); lean_dec_ref(x_3); -x_154 = l_Lean_IR_VarId_alphaEqv(x_1, x_150, x_152); -lean_dec(x_152); -lean_dec(x_150); -if (x_154 == 0) -{ +x_157 = l_Lean_IR_VarId_alphaEqv(x_1, x_153, x_155); +lean_dec(x_155); lean_dec(x_153); -lean_dec(x_151); +if (x_157 == 0) +{ +lean_dec(x_156); +lean_dec(x_154); lean_dec(x_1); -return x_154; +return x_157; } else { -x_2 = x_151; -x_3 = x_153; +x_2 = x_154; +x_3 = x_156; goto _start; } } else { -uint8_t x_156; +uint8_t x_159; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_156 = 0; -return x_156; +x_159 = 0; +return x_159; } } case 9: { if (lean_obj_tag(x_3) == 9) { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; uint8_t x_169; -x_157 = lean_ctor_get(x_2, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_2, 1); -lean_inc(x_158); -x_159 = lean_ctor_get(x_2, 3); -lean_inc_ref(x_159); -lean_dec_ref(x_2); -x_160 = lean_ctor_get(x_3, 0); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; uint8_t x_172; +x_160 = lean_ctor_get(x_2, 0); lean_inc(x_160); -x_161 = lean_ctor_get(x_3, 1); +x_161 = lean_ctor_get(x_2, 1); lean_inc(x_161); -x_162 = lean_ctor_get(x_3, 3); +x_162 = lean_ctor_get(x_2, 3); lean_inc_ref(x_162); +lean_dec_ref(x_2); +x_163 = lean_ctor_get(x_3, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_3, 1); +lean_inc(x_164); +x_165 = lean_ctor_get(x_3, 3); +lean_inc_ref(x_165); lean_dec_ref(x_3); -x_169 = lean_name_eq(x_157, x_160); +x_172 = lean_name_eq(x_160, x_163); +lean_dec(x_163); lean_dec(x_160); -lean_dec(x_157); -if (x_169 == 0) +if (x_172 == 0) { +lean_dec(x_164); lean_dec(x_161); -lean_dec(x_158); -x_163 = x_169; -goto block_168; +x_166 = x_172; +goto block_171; } else { -uint8_t x_170; -x_170 = l_Lean_IR_VarId_alphaEqv(x_1, x_158, x_161); +uint8_t x_173; +x_173 = l_Lean_IR_VarId_alphaEqv(x_1, x_161, x_164); +lean_dec(x_164); lean_dec(x_161); -lean_dec(x_158); -x_163 = x_170; -goto block_168; +x_166 = x_173; +goto block_171; } -block_168: +block_171: { -if (x_163 == 0) +if (x_166 == 0) { +lean_dec_ref(x_165); lean_dec_ref(x_162); -lean_dec_ref(x_159); lean_dec(x_1); -return x_163; +return x_166; } else { -lean_object* x_164; lean_object* x_165; uint8_t x_166; -x_164 = lean_array_get_size(x_159); -x_165 = lean_array_get_size(x_162); -x_166 = lean_nat_dec_eq(x_164, x_165); -lean_dec(x_165); -if (x_166 == 0) +lean_object* x_167; lean_object* x_168; uint8_t x_169; +x_167 = lean_array_get_size(x_162); +x_168 = lean_array_get_size(x_165); +x_169 = lean_nat_dec_eq(x_167, x_168); +lean_dec(x_168); +if (x_169 == 0) { -lean_dec(x_164); +lean_dec(x_167); +lean_dec_ref(x_165); lean_dec_ref(x_162); -lean_dec_ref(x_159); lean_dec(x_1); -return x_166; +return x_169; } else { -uint8_t x_167; -x_167 = l_Array_isEqvAux___at___00Lean_IR_FnBody_alphaEqv_spec__0___redArg(x_1, x_159, x_162, x_164); +uint8_t x_170; +x_170 = l_Array_isEqvAux___at___00Lean_IR_FnBody_alphaEqv_spec__0___redArg(x_1, x_162, x_165, x_167); +lean_dec_ref(x_165); lean_dec_ref(x_162); -lean_dec_ref(x_159); -return x_167; +return x_170; } } } } else { -uint8_t x_171; +uint8_t x_174; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_171 = 0; -return x_171; +x_174 = 0; +return x_174; } } case 10: { if (lean_obj_tag(x_3) == 10) { -lean_object* x_172; lean_object* x_173; uint8_t x_174; -x_172 = lean_ctor_get(x_2, 0); -lean_inc(x_172); +lean_object* x_175; lean_object* x_176; uint8_t x_177; +x_175 = lean_ctor_get(x_2, 0); +lean_inc(x_175); lean_dec_ref(x_2); -x_173 = lean_ctor_get(x_3, 0); -lean_inc(x_173); +x_176 = lean_ctor_get(x_3, 0); +lean_inc(x_176); lean_dec_ref(x_3); -x_174 = l_Lean_IR_Arg_alphaEqv(x_1, x_172, x_173); -lean_dec(x_173); -lean_dec(x_172); +x_177 = l_Lean_IR_Arg_alphaEqv(x_1, x_175, x_176); +lean_dec(x_176); +lean_dec(x_175); lean_dec(x_1); -return x_174; +return x_177; } else { -uint8_t x_175; +uint8_t x_178; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_175 = 0; -return x_175; +x_178 = 0; +return x_178; } } case 11: { if (lean_obj_tag(x_3) == 11) { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_176 = lean_ctor_get(x_2, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_177); +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; +x_179 = lean_ctor_get(x_2, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_180); lean_dec_ref(x_2); -x_178 = lean_ctor_get(x_3, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_179); +x_181 = lean_ctor_get(x_3, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_182); lean_dec_ref(x_3); -x_180 = lean_nat_dec_eq(x_176, x_178); -lean_dec(x_178); -lean_dec(x_176); -if (x_180 == 0) +x_183 = lean_nat_dec_eq(x_179, x_181); +lean_dec(x_181); +lean_dec(x_179); +if (x_183 == 0) { -lean_dec_ref(x_179); -lean_dec_ref(x_177); +lean_dec_ref(x_182); +lean_dec_ref(x_180); lean_dec(x_1); -return x_180; +return x_183; } else { -uint8_t x_181; -x_181 = l_Lean_IR_args_alphaEqv(x_1, x_177, x_179); -lean_dec_ref(x_179); -lean_dec_ref(x_177); +uint8_t x_184; +x_184 = l_Lean_IR_args_alphaEqv(x_1, x_180, x_182); +lean_dec_ref(x_182); +lean_dec_ref(x_180); lean_dec(x_1); -return x_181; +return x_184; } } else { -uint8_t x_182; +uint8_t x_185; lean_dec(x_3); lean_dec_ref(x_2); lean_dec(x_1); -x_182 = 0; -return x_182; +x_185 = 0; +return x_185; } } default: @@ -14222,52 +14958,52 @@ return x_182; lean_dec(x_1); if (lean_obj_tag(x_3) == 12) { -uint8_t x_183; -x_183 = 1; -return x_183; +uint8_t x_186; +x_186 = 1; +return x_186; } else { -uint8_t x_184; +uint8_t x_187; lean_dec(x_3); -x_184 = 0; -return x_184; +x_187 = 0; +return x_187; } } } block_11: { -if (x_4 == 0) +if (x_6 == 0) { -if (x_7 == 0) +if (x_5 == 0) { -x_1 = x_8; -x_2 = x_5; -x_3 = x_6; +x_1 = x_7; +x_2 = x_8; +x_3 = x_4; goto _start; } else { lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -return x_4; +lean_dec(x_7); +lean_dec(x_4); +return x_6; } } else { -if (x_7 == 0) +if (x_5 == 0) { lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -return x_7; +lean_dec(x_7); +lean_dec(x_4); +return x_5; } else { -x_1 = x_8; -x_2 = x_5; -x_3 = x_6; +x_1 = x_7; +x_2 = x_8; +x_3 = x_4; goto _start; } } @@ -14276,48 +15012,48 @@ goto _start; { if (x_19 == 0) { -lean_dec(x_17); -lean_dec(x_14); +lean_dec(x_18); +lean_dec(x_15); lean_dec(x_13); return x_19; } else { -if (x_18 == 0) +if (x_17 == 0) { if (x_16 == 0) { -x_4 = x_12; -x_5 = x_13; +x_4 = x_13; +x_5 = x_12; x_6 = x_14; x_7 = x_15; -x_8 = x_17; +x_8 = x_18; goto block_11; } else { -lean_dec(x_17); -lean_dec(x_14); +lean_dec(x_18); +lean_dec(x_15); lean_dec(x_13); -return x_18; +return x_17; } } else { if (x_16 == 0) { -lean_dec(x_17); -lean_dec(x_14); +lean_dec(x_18); +lean_dec(x_15); lean_dec(x_13); return x_16; } else { -x_4 = x_12; -x_5 = x_13; +x_4 = x_13; +x_5 = x_12; x_6 = x_14; x_7 = x_15; -x_8 = x_17; +x_8 = x_18; goto block_11; } } @@ -14333,13 +15069,13 @@ if (x_32 == 0) { lean_dec(x_28); lean_dec(x_23); -x_12 = x_25; -x_13 = x_26; -x_14 = x_31; -x_15 = x_30; +x_12 = x_30; +x_13 = x_31; +x_14 = x_25; +x_15 = x_21; x_16 = x_29; -x_17 = x_21; -x_18 = x_24; +x_17 = x_24; +x_18 = x_26; x_19 = x_32; goto block_20; } @@ -14349,13 +15085,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_26; -x_14 = x_31; -x_15 = x_30; +x_12 = x_30; +x_13 = x_31; +x_14 = x_25; +x_15 = x_21; x_16 = x_29; -x_17 = x_21; -x_18 = x_24; +x_17 = x_24; +x_18 = x_26; x_19 = x_33; goto block_20; } @@ -14589,6 +15325,30 @@ static lean_object* _init_l_Lean_IR_getUnboxOpName___closed__5() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_int32", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_getUnboxOpName___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_int64", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_getUnboxOpName___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_isize", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_getUnboxOpName___closed__8() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("lean_unbox", 10, 10); return x_1; } @@ -14627,12 +15387,30 @@ lean_object* x_6; x_6 = l_Lean_IR_getUnboxOpName___closed__4; return x_6; } -default: +case 16: { lean_object* x_7; x_7 = l_Lean_IR_getUnboxOpName___closed__5; return x_7; } +case 17: +{ +lean_object* x_8; +x_8 = l_Lean_IR_getUnboxOpName___closed__6; +return x_8; +} +case 18: +{ +lean_object* x_9; +x_9 = l_Lean_IR_getUnboxOpName___closed__7; +return x_9; +} +default: +{ +lean_object* x_10; +x_10 = l_Lean_IR_getUnboxOpName___closed__8; +return x_10; +} } } } @@ -14822,6 +15600,26 @@ l_Lean_IR_instReprIRType_repr___closed__30 = _init_l_Lean_IR_instReprIRType_repr lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__30); l_Lean_IR_instReprIRType_repr___closed__31 = _init_l_Lean_IR_instReprIRType_repr___closed__31(); lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__31); +l_Lean_IR_instReprIRType_repr___closed__32 = _init_l_Lean_IR_instReprIRType_repr___closed__32(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__32); +l_Lean_IR_instReprIRType_repr___closed__33 = _init_l_Lean_IR_instReprIRType_repr___closed__33(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__33); +l_Lean_IR_instReprIRType_repr___closed__34 = _init_l_Lean_IR_instReprIRType_repr___closed__34(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__34); +l_Lean_IR_instReprIRType_repr___closed__35 = _init_l_Lean_IR_instReprIRType_repr___closed__35(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__35); +l_Lean_IR_instReprIRType_repr___closed__36 = _init_l_Lean_IR_instReprIRType_repr___closed__36(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__36); +l_Lean_IR_instReprIRType_repr___closed__37 = _init_l_Lean_IR_instReprIRType_repr___closed__37(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__37); +l_Lean_IR_instReprIRType_repr___closed__38 = _init_l_Lean_IR_instReprIRType_repr___closed__38(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__38); +l_Lean_IR_instReprIRType_repr___closed__39 = _init_l_Lean_IR_instReprIRType_repr___closed__39(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__39); +l_Lean_IR_instReprIRType_repr___closed__40 = _init_l_Lean_IR_instReprIRType_repr___closed__40(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__40); +l_Lean_IR_instReprIRType_repr___closed__41 = _init_l_Lean_IR_instReprIRType_repr___closed__41(); +lean_mark_persistent(l_Lean_IR_instReprIRType_repr___closed__41); l_Lean_IR_instReprIRType___closed__0 = _init_l_Lean_IR_instReprIRType___closed__0(); lean_mark_persistent(l_Lean_IR_instReprIRType___closed__0); l_Lean_IR_instReprIRType = _init_l_Lean_IR_instReprIRType(); @@ -15056,6 +15854,12 @@ l_Lean_IR_getUnboxOpName___closed__4 = _init_l_Lean_IR_getUnboxOpName___closed__ lean_mark_persistent(l_Lean_IR_getUnboxOpName___closed__4); l_Lean_IR_getUnboxOpName___closed__5 = _init_l_Lean_IR_getUnboxOpName___closed__5(); lean_mark_persistent(l_Lean_IR_getUnboxOpName___closed__5); +l_Lean_IR_getUnboxOpName___closed__6 = _init_l_Lean_IR_getUnboxOpName___closed__6(); +lean_mark_persistent(l_Lean_IR_getUnboxOpName___closed__6); +l_Lean_IR_getUnboxOpName___closed__7 = _init_l_Lean_IR_getUnboxOpName___closed__7(); +lean_mark_persistent(l_Lean_IR_getUnboxOpName___closed__7); +l_Lean_IR_getUnboxOpName___closed__8 = _init_l_Lean_IR_getUnboxOpName___closed__8(); +lean_mark_persistent(l_Lean_IR_getUnboxOpName___closed__8); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/IR/Boxing.c b/stage0/stdlib/Lean/Compiler/IR/Boxing.c index 99128e8cb865..96cd5ba4e2b0 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Boxing.c +++ b/stage0/stdlib/Lean/Compiler/IR/Boxing.c @@ -118,11 +118,11 @@ static lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___at___0 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*); +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*); 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_object* l_Lean_Name_str___override(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*, uint8_t, 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*); @@ -133,6 +133,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_castVarIfNeeded(lean_object*, LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_isBoxedName___boxed(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_object* l_Lean_IR_IRType_ptrSizedTypeForSign(uint8_t); LEAN_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody(lean_object*, lean_object*, lean_object*); 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*); @@ -4070,19 +4071,20 @@ goto _start; } } } -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_EXPORT lean_object* l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(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_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_8; lean_object* x_9; +x_8 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_5); +lean_ctor_set(x_8, 3, x_4); +lean_ctor_set_uint8(x_8, sizeof(void*)*4, x_3); +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) { @@ -4430,15 +4432,17 @@ return x_85; } case 4: { -uint8_t x_86; -x_86 = !lean_is_exclusive(x_1); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_87 = lean_ctor_get(x_1, 0); -x_88 = lean_ctor_get(x_1, 1); -x_89 = lean_ctor_get(x_1, 2); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_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_uint8(x_1, sizeof(void*)*4); x_90 = lean_ctor_get(x_1, 3); +lean_inc(x_90); +lean_dec_ref(x_1); lean_inc_ref(x_2); x_91 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_90, x_2, x_3); x_92 = lean_ctor_get(x_91, 0); @@ -4446,13 +4450,14 @@ 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_ExplicitBoxing_getVarType(x_89, x_2, x_93); +x_94 = l_Lean_IR_ExplicitBoxing_getVarType(x_88, x_2, x_93); x_95 = lean_ctor_get(x_94, 0); lean_inc(x_95); x_96 = lean_ctor_get(x_94, 1); lean_inc(x_96); lean_dec_ref(x_94); -x_97 = lean_box(5); +x_97 = l_Lean_IR_IRType_ptrSizedTypeForSign(x_89); +lean_inc(x_97); lean_inc(x_95); x_98 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_95, x_97); if (x_98 == 0) @@ -4465,359 +4470,362 @@ x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); lean_dec_ref(x_99); lean_inc_ref(x_2); -x_102 = l_Lean_IR_ExplicitBoxing_mkCast(x_89, x_95, x_97, x_2, x_101); +lean_inc(x_97); +x_102 = l_Lean_IR_ExplicitBoxing_mkCast(x_88, x_95, x_97, x_2, x_101); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); x_104 = lean_ctor_get(x_102, 1); lean_inc(x_104); lean_dec_ref(x_102); lean_inc(x_100); -x_105 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_87, x_88, x_92, x_100, x_2, x_104); +x_105 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_86, x_87, x_89, x_92, x_100, x_2, x_104); lean_dec_ref(x_2); x_106 = !lean_is_exclusive(x_105); if (x_106 == 0) { -lean_object* x_107; +lean_object* x_107; lean_object* x_108; x_107 = lean_ctor_get(x_105, 0); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_107); -lean_ctor_set(x_1, 2, x_103); -lean_ctor_set(x_1, 1, x_97); -lean_ctor_set(x_1, 0, x_100); -lean_ctor_set(x_105, 0, x_1); +x_108 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_108, 0, x_100); +lean_ctor_set(x_108, 1, x_97); +lean_ctor_set(x_108, 2, x_103); +lean_ctor_set(x_108, 3, x_107); +lean_ctor_set(x_105, 0, x_108); return x_105; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_105, 0); -x_109 = lean_ctor_get(x_105, 1); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_109 = lean_ctor_get(x_105, 0); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); lean_inc(x_109); -lean_inc(x_108); lean_dec(x_105); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_108); -lean_ctor_set(x_1, 2, x_103); -lean_ctor_set(x_1, 1, x_97); -lean_ctor_set(x_1, 0, x_100); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_1); -lean_ctor_set(x_110, 1, x_109); -return x_110; +x_111 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_111, 0, x_100); +lean_ctor_set(x_111, 1, x_97); +lean_ctor_set(x_111, 2, x_103); +lean_ctor_set(x_111, 3, x_109); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +return x_112; } } else { -lean_object* x_111; +lean_object* x_113; +lean_dec(x_97); lean_dec(x_95); -lean_free_object(x_1); -x_111 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_87, x_88, x_92, x_89, x_2, x_96); +x_113 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_86, x_87, x_89, x_92, x_88, x_2, x_96); lean_dec_ref(x_2); -return x_111; +return x_113; } } -else +case 5: { -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; uint8_t x_123; -x_112 = lean_ctor_get(x_1, 0); -x_113 = lean_ctor_get(x_1, 1); -x_114 = lean_ctor_get(x_1, 2); -x_115 = lean_ctor_get(x_1, 3); -lean_inc(x_115); +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; uint8_t x_126; +x_114 = lean_ctor_get(x_1, 0); lean_inc(x_114); -lean_inc(x_113); -lean_inc(x_112); -lean_dec(x_1); -lean_inc_ref(x_2); -x_116 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_115, x_2, x_3); -x_117 = lean_ctor_get(x_116, 0); +x_115 = lean_ctor_get(x_1, 1); +lean_inc(x_115); +x_116 = lean_ctor_get(x_1, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_1, 3); lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); +x_118 = lean_ctor_get(x_1, 4); lean_inc(x_118); -lean_dec_ref(x_116); -x_119 = l_Lean_IR_ExplicitBoxing_getVarType(x_114, x_2, x_118); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_119, 1); +x_119 = lean_ctor_get(x_1, 5); +lean_inc(x_119); +lean_dec_ref(x_1); +lean_inc_ref(x_2); +x_120 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_119, x_2, x_3); +x_121 = lean_ctor_get(x_120, 0); lean_inc(x_121); -lean_dec_ref(x_119); -x_122 = lean_box(5); -lean_inc(x_120); -x_123 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_120, x_122); -if (x_123 == 0) -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_124 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_121); -x_125 = lean_ctor_get(x_124, 0); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec_ref(x_120); +x_123 = l_Lean_IR_ExplicitBoxing_getVarType(x_117, x_2, x_122); +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec_ref(x_124); -lean_inc_ref(x_2); -x_127 = l_Lean_IR_ExplicitBoxing_mkCast(x_114, x_120, x_122, x_2, x_126); +lean_dec_ref(x_123); +lean_inc(x_118); +lean_inc(x_124); +x_126 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_124, x_118); +if (x_126 == 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; uint8_t x_134; +x_127 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_125); x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); x_129 = lean_ctor_get(x_127, 1); lean_inc(x_129); lean_dec_ref(x_127); -lean_inc(x_125); -x_130 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_112, x_113, x_117, x_125, x_2, x_129); -lean_dec_ref(x_2); +lean_inc_ref(x_2); +lean_inc(x_118); +x_130 = l_Lean_IR_ExplicitBoxing_mkCast(x_117, x_124, x_118, 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); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_133 = x_130; -} else { - lean_dec_ref(x_130); - x_133 = lean_box(0); -} -x_134 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_134, 0, x_125); -lean_ctor_set(x_134, 1, x_122); -lean_ctor_set(x_134, 2, x_128); -lean_ctor_set(x_134, 3, x_131); -if (lean_is_scalar(x_133)) { - x_135 = lean_alloc_ctor(0, 2, 0); -} else { - x_135 = x_133; +lean_dec_ref(x_130); +lean_inc(x_128); +lean_inc(x_118); +x_133 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_114, x_115, x_116, x_118, x_121, x_128, x_2, x_132); +lean_dec_ref(x_2); +x_134 = !lean_is_exclusive(x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; +x_135 = lean_ctor_get(x_133, 0); +x_136 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_136, 0, x_128); +lean_ctor_set(x_136, 1, x_118); +lean_ctor_set(x_136, 2, x_131); +lean_ctor_set(x_136, 3, x_135); +lean_ctor_set(x_133, 0, x_136); +return x_133; +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_137 = lean_ctor_get(x_133, 0); +x_138 = lean_ctor_get(x_133, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_133); +x_139 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_139, 0, x_128); +lean_ctor_set(x_139, 1, x_118); +lean_ctor_set(x_139, 2, x_131); +lean_ctor_set(x_139, 3, x_137); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_138); +return x_140; } -lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_132); -return x_135; } else { -lean_object* x_136; -lean_dec(x_120); -x_136 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_112, x_113, x_117, x_114, x_2, x_121); +lean_object* x_141; +lean_dec(x_124); +x_141 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_114, x_115, x_116, x_118, x_121, x_117, x_2, x_125); lean_dec_ref(x_2); -return x_136; -} +return x_141; } } -case 5: +case 9: { -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; -x_137 = lean_ctor_get(x_1, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_1, 1); -lean_inc(x_138); -x_139 = lean_ctor_get(x_1, 2); -lean_inc(x_139); -x_140 = lean_ctor_get(x_1, 3); -lean_inc(x_140); -x_141 = lean_ctor_get(x_1, 4); -lean_inc(x_141); -x_142 = lean_ctor_get(x_1, 5); -lean_inc(x_142); -lean_dec_ref(x_1); +uint8_t x_142; +x_142 = !lean_is_exclusive(x_1); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; size_t x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_143 = lean_ctor_get(x_1, 0); +x_144 = lean_ctor_get(x_1, 1); +x_145 = lean_ctor_get(x_1, 2); +x_146 = lean_ctor_get(x_1, 3); +x_147 = lean_array_size(x_146); +x_148 = 0; lean_inc_ref(x_2); -x_143 = l_Lean_IR_ExplicitBoxing_visitFnBody(x_142, x_2, 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_ExplicitBoxing_getVarType(x_140, x_2, x_145); -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); -lean_inc(x_141); -lean_inc(x_147); -x_149 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_147, x_141); -if (x_149 == 0) -{ -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_150 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_148); -x_151 = lean_ctor_get(x_150, 0); +x_149 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_147, x_148, x_146, x_2, x_3); +x_150 = lean_ctor_get(x_149, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 1); lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec_ref(x_150); -lean_inc_ref(x_2); -lean_inc(x_141); -x_153 = l_Lean_IR_ExplicitBoxing_mkCast(x_140, x_147, x_141, x_2, x_152); -x_154 = lean_ctor_get(x_153, 0); +lean_dec_ref(x_149); +x_152 = l_Lean_IR_ExplicitBoxing_getVarType(x_144, x_2, x_151); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); -lean_inc(x_155); -lean_dec_ref(x_153); -lean_inc(x_151); -lean_inc(x_141); -x_156 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_137, x_138, x_139, x_141, x_144, x_151, x_2, x_155); +lean_dec_ref(x_152); +lean_inc(x_145); +lean_inc(x_153); +x_155 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_153, x_145); +if (x_155 == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_156 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_154); +x_157 = lean_ctor_get(x_156, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_156, 1); +lean_inc(x_158); +lean_dec_ref(x_156); +lean_inc_ref(x_2); +lean_inc(x_145); +x_159 = l_Lean_IR_ExplicitBoxing_mkCast(x_144, x_153, x_145, x_2, x_158); +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); +lean_inc(x_157); +lean_inc(x_145); +x_162 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_143, x_145, x_150, x_157, x_2, x_161); lean_dec_ref(x_2); -x_157 = !lean_is_exclusive(x_156); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; -x_158 = lean_ctor_get(x_156, 0); -x_159 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_159, 0, x_151); -lean_ctor_set(x_159, 1, x_141); -lean_ctor_set(x_159, 2, x_154); -lean_ctor_set(x_159, 3, x_158); -lean_ctor_set(x_156, 0, x_159); -return x_156; +x_163 = !lean_is_exclusive(x_162); +if (x_163 == 0) +{ +lean_object* x_164; +x_164 = lean_ctor_get(x_162, 0); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 3, x_164); +lean_ctor_set(x_1, 2, x_160); +lean_ctor_set(x_1, 1, x_145); +lean_ctor_set(x_1, 0, x_157); +lean_ctor_set(x_162, 0, x_1); +return x_162; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_160 = lean_ctor_get(x_156, 0); -x_161 = lean_ctor_get(x_156, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_156); -x_162 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_162, 0, x_151); -lean_ctor_set(x_162, 1, x_141); -lean_ctor_set(x_162, 2, x_154); -lean_ctor_set(x_162, 3, x_160); -x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_161); -return x_163; +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_162, 0); +x_166 = lean_ctor_get(x_162, 1); +lean_inc(x_166); +lean_inc(x_165); +lean_dec(x_162); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 3, x_165); +lean_ctor_set(x_1, 2, x_160); +lean_ctor_set(x_1, 1, x_145); +lean_ctor_set(x_1, 0, x_157); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_1); +lean_ctor_set(x_167, 1, x_166); +return x_167; } } else { -lean_object* x_164; -lean_dec(x_147); -x_164 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__1(x_137, x_138, x_139, x_141, x_144, x_140, x_2, x_148); +lean_object* x_168; +lean_dec(x_153); +lean_free_object(x_1); +x_168 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_143, x_145, x_150, x_144, x_2, x_154); lean_dec_ref(x_2); -return x_164; +return x_168; } } -case 9: +else { -uint8_t x_165; -x_165 = !lean_is_exclusive(x_1); -if (x_165 == 0) -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; size_t x_170; size_t x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_166 = lean_ctor_get(x_1, 0); -x_167 = lean_ctor_get(x_1, 1); -x_168 = lean_ctor_get(x_1, 2); -x_169 = lean_ctor_get(x_1, 3); -x_170 = lean_array_size(x_169); -x_171 = 0; +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); +lean_inc(x_172); +lean_inc(x_171); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_1); +x_173 = lean_array_size(x_172); +x_174 = 0; lean_inc_ref(x_2); -x_172 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_170, x_171, x_169, x_2, x_3); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec_ref(x_172); -x_175 = l_Lean_IR_ExplicitBoxing_getVarType(x_167, x_2, x_174); +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); -lean_inc(x_168); -lean_inc(x_176); -x_178 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_176, x_168); -if (x_178 == 0) -{ -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; -x_179 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_177); -x_180 = lean_ctor_get(x_179, 0); +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); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); -lean_dec_ref(x_179); -lean_inc_ref(x_2); -lean_inc(x_168); -x_182 = l_Lean_IR_ExplicitBoxing_mkCast(x_167, x_176, x_168, x_2, x_181); +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_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; +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(x_180); -lean_inc(x_168); -x_185 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_166, x_168, x_173, x_180, x_2, x_184); +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_dec_ref(x_2); -x_186 = !lean_is_exclusive(x_185); -if (x_186 == 0) -{ -lean_object* x_187; -x_187 = lean_ctor_get(x_185, 0); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_187); -lean_ctor_set(x_1, 2, x_183); -lean_ctor_set(x_1, 1, x_168); -lean_ctor_set(x_1, 0, x_180); -lean_ctor_set(x_185, 0, x_1); -return x_185; -} -else -{ -lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_188 = lean_ctor_get(x_185, 0); -x_189 = lean_ctor_get(x_185, 1); +x_189 = lean_ctor_get(x_188, 0); lean_inc(x_189); -lean_inc(x_188); -lean_dec(x_185); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 3, x_188); -lean_ctor_set(x_1, 2, x_183); -lean_ctor_set(x_1, 1, x_168); -lean_ctor_set(x_1, 0, x_180); -x_190 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_190, 0, x_1); -lean_ctor_set(x_190, 1, x_189); -return x_190; +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; +} else { + lean_dec_ref(x_188); + x_191 = lean_box(0); +} +x_192 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_192, 0, x_183); +lean_ctor_set(x_192, 1, x_171); +lean_ctor_set(x_192, 2, x_186); +lean_ctor_set(x_192, 3, x_189); +if (lean_is_scalar(x_191)) { + x_193 = lean_alloc_ctor(0, 2, 0); +} else { + x_193 = x_191; } +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_190); +return x_193; } else { -lean_object* x_191; -lean_dec(x_176); -lean_free_object(x_1); -x_191 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_166, x_168, x_173, x_167, x_2, x_177); +lean_object* x_194; +lean_dec(x_179); +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_191; +return x_194; } } -else +} +case 10: { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; size_t x_196; size_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; -x_192 = lean_ctor_get(x_1, 0); -x_193 = lean_ctor_get(x_1, 1); -x_194 = lean_ctor_get(x_1, 2); -x_195 = lean_ctor_get(x_1, 3); +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_195 = lean_ctor_get(x_1, 0); lean_inc(x_195); -lean_inc(x_194); -lean_inc(x_193); -lean_inc(x_192); -lean_dec(x_1); -x_196 = lean_array_size(x_195); -x_197 = 0; -lean_inc_ref(x_2); -x_198 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitBoxing_visitFnBody_spec__0(x_196, x_197, x_195, x_2, x_3); -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); +lean_dec_ref(x_1); +x_196 = l_Lean_IR_ExplicitBoxing_getResultType(x_2, x_3); +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); +x_199 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed), 3, 0); +if (lean_obj_tag(x_195) == 0) +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +x_200 = lean_ctor_get(x_195, 0); lean_inc(x_200); -lean_dec_ref(x_198); -x_201 = l_Lean_IR_ExplicitBoxing_getVarType(x_193, x_2, x_200); +lean_dec_ref(x_195); +x_201 = l_Lean_IR_ExplicitBoxing_getVarType(x_200, x_2, x_198); 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); -lean_inc(x_194); +lean_inc(x_197); lean_inc(x_202); -x_204 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_202, x_194); +x_204 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_202, x_197); if (x_204 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_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_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; uint8_t x_212; x_205 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_203); x_206 = lean_ctor_get(x_205, 0); lean_inc(x_206); @@ -4825,264 +4833,179 @@ x_207 = lean_ctor_get(x_205, 1); lean_inc(x_207); lean_dec_ref(x_205); lean_inc_ref(x_2); -lean_inc(x_194); -x_208 = l_Lean_IR_ExplicitBoxing_mkCast(x_193, x_202, x_194, x_2, x_207); +lean_inc(x_197); +x_208 = l_Lean_IR_ExplicitBoxing_mkCast(x_200, x_202, x_197, x_2, x_207); 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(x_206); -lean_inc(x_194); -x_211 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_192, x_194, x_199, x_206, x_2, x_210); -lean_dec_ref(x_2); -x_212 = lean_ctor_get(x_211, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_211, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_211)) { - lean_ctor_release(x_211, 0); - lean_ctor_release(x_211, 1); - x_214 = x_211; -} else { - lean_dec_ref(x_211); - x_214 = lean_box(0); -} -x_215 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_215, 0, x_206); -lean_ctor_set(x_215, 1, x_194); -lean_ctor_set(x_215, 2, x_209); -lean_ctor_set(x_215, 3, x_212); -if (lean_is_scalar(x_214)) { - x_216 = lean_alloc_ctor(0, 2, 0); -} else { - x_216 = x_214; +x_211 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_199, x_206, x_2, x_210); +x_212 = !lean_is_exclusive(x_211); +if (x_212 == 0) +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_ctor_get(x_211, 0); +x_214 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_214, 0, x_206); +lean_ctor_set(x_214, 1, x_197); +lean_ctor_set(x_214, 2, x_209); +lean_ctor_set(x_214, 3, x_213); +lean_ctor_set(x_211, 0, x_214); +return x_211; +} +else +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_215 = lean_ctor_get(x_211, 0); +x_216 = lean_ctor_get(x_211, 1); +lean_inc(x_216); +lean_inc(x_215); +lean_dec(x_211); +x_217 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_217, 0, x_206); +lean_ctor_set(x_217, 1, x_197); +lean_ctor_set(x_217, 2, x_209); +lean_ctor_set(x_217, 3, x_215); +x_218 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_218, 0, x_217); +lean_ctor_set(x_218, 1, x_216); +return x_218; } -lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_213); -return x_216; } else { -lean_object* x_217; +lean_object* x_219; lean_dec(x_202); -x_217 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__2(x_192, x_194, x_199, x_193, x_2, x_203); -lean_dec_ref(x_2); -return x_217; +lean_dec(x_197); +x_219 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_199, x_200, x_2, x_203); +return x_219; } } +else +{ +lean_object* x_220; +lean_dec_ref(x_199); +lean_dec(x_197); +x_220 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_195, x_2, x_198); +lean_dec_ref(x_2); +return x_220; } -case 10: +} +case 11: { -lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_218 = lean_ctor_get(x_1, 0); -lean_inc(x_218); -lean_dec_ref(x_1); -x_219 = l_Lean_IR_ExplicitBoxing_getResultType(x_2, x_3); -x_220 = lean_ctor_get(x_219, 0); -lean_inc(x_220); -x_221 = lean_ctor_get(x_219, 1); -lean_inc(x_221); -lean_dec_ref(x_219); -x_222 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3___boxed), 3, 0); -if (lean_obj_tag(x_218) == 0) -{ -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; uint8_t x_227; -x_223 = lean_ctor_get(x_218, 0); -lean_inc(x_223); -lean_dec_ref(x_218); -x_224 = l_Lean_IR_ExplicitBoxing_getVarType(x_223, x_2, x_221); +uint8_t x_221; +x_221 = !lean_is_exclusive(x_1); +if (x_221 == 0) +{ +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; uint8_t x_232; +x_222 = lean_ctor_get(x_1, 0); +x_223 = lean_ctor_get(x_1, 1); +x_224 = l_Lean_IR_ExplicitBoxing_getJPParams(x_222, x_2, x_3); x_225 = lean_ctor_get(x_224, 0); lean_inc(x_225); x_226 = lean_ctor_get(x_224, 1); lean_inc(x_226); lean_dec_ref(x_224); -lean_inc(x_220); -lean_inc(x_225); -x_227 = l_Lean_IR_ExplicitBoxing_eqvTypes(x_225, x_220); -if (x_227 == 0) -{ -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; uint8_t x_235; -x_228 = l___private_Lean_Compiler_IR_Boxing_0__Lean_IR_ExplicitBoxing_M_mkFresh___redArg(x_226); -x_229 = lean_ctor_get(x_228, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_228, 1); +x_227 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_228 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); +lean_closure_set(x_228, 0, x_227); +lean_closure_set(x_228, 1, x_225); +x_229 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_223, x_228, x_2, x_226); +lean_dec_ref(x_223); +x_230 = lean_ctor_get(x_229, 0); lean_inc(x_230); -lean_dec_ref(x_228); -lean_inc_ref(x_2); -lean_inc(x_220); -x_231 = l_Lean_IR_ExplicitBoxing_mkCast(x_223, x_225, x_220, x_2, x_230); -x_232 = lean_ctor_get(x_231, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_231, 1); -lean_inc(x_233); -lean_dec_ref(x_231); -lean_inc(x_229); -x_234 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_222, x_229, x_2, x_233); -x_235 = !lean_is_exclusive(x_234); -if (x_235 == 0) -{ -lean_object* x_236; lean_object* x_237; -x_236 = lean_ctor_get(x_234, 0); -x_237 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_237, 0, x_229); -lean_ctor_set(x_237, 1, x_220); -lean_ctor_set(x_237, 2, x_232); -lean_ctor_set(x_237, 3, x_236); -lean_ctor_set(x_234, 0, x_237); -return x_234; +x_231 = lean_ctor_get(x_229, 1); +lean_inc(x_231); +lean_dec_ref(x_229); +x_232 = !lean_is_exclusive(x_230); +if (x_232 == 0) +{ +lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_233 = lean_ctor_get(x_230, 0); +x_234 = lean_ctor_get(x_230, 1); +lean_ctor_set(x_1, 1, x_233); +x_235 = l_Lean_IR_reshape(x_234, x_1); +lean_ctor_set(x_230, 1, x_231); +lean_ctor_set(x_230, 0, x_235); +return x_230; } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_238 = lean_ctor_get(x_234, 0); -x_239 = lean_ctor_get(x_234, 1); -lean_inc(x_239); -lean_inc(x_238); -lean_dec(x_234); -x_240 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_240, 0, x_229); -lean_ctor_set(x_240, 1, x_220); -lean_ctor_set(x_240, 2, x_232); -lean_ctor_set(x_240, 3, x_238); -x_241 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_241, 0, x_240); -lean_ctor_set(x_241, 1, x_239); -return x_241; +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_236 = lean_ctor_get(x_230, 0); +x_237 = lean_ctor_get(x_230, 1); +lean_inc(x_237); +lean_inc(x_236); +lean_dec(x_230); +lean_ctor_set(x_1, 1, x_236); +x_238 = l_Lean_IR_reshape(x_237, x_1); +x_239 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_239, 0, x_238); +lean_ctor_set(x_239, 1, x_231); +return x_239; } } else { -lean_object* x_242; -lean_dec(x_225); -lean_dec(x_220); -x_242 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__4(x_222, x_223, x_2, x_226); -return x_242; -} -} -else -{ -lean_object* x_243; -lean_dec_ref(x_222); -lean_dec(x_220); -x_243 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__3(x_218, x_2, x_221); -lean_dec_ref(x_2); -return x_243; -} -} -case 11: -{ -uint8_t x_244; -x_244 = !lean_is_exclusive(x_1); -if (x_244 == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; -x_245 = lean_ctor_get(x_1, 0); -x_246 = lean_ctor_get(x_1, 1); -x_247 = l_Lean_IR_ExplicitBoxing_getJPParams(x_245, x_2, x_3); +lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +x_240 = lean_ctor_get(x_1, 0); +x_241 = lean_ctor_get(x_1, 1); +lean_inc(x_241); +lean_inc(x_240); +lean_dec(x_1); +x_242 = l_Lean_IR_ExplicitBoxing_getJPParams(x_240, x_2, x_3); +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); +lean_inc(x_244); +lean_dec_ref(x_242); +x_245 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; +x_246 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); +lean_closure_set(x_246, 0, x_245); +lean_closure_set(x_246, 1, x_243); +x_247 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_241, x_246, x_2, x_244); +lean_dec_ref(x_241); x_248 = lean_ctor_get(x_247, 0); lean_inc(x_248); x_249 = lean_ctor_get(x_247, 1); lean_inc(x_249); lean_dec_ref(x_247); -x_250 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_251 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); -lean_closure_set(x_251, 0, x_250); -lean_closure_set(x_251, 1, x_248); -x_252 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_246, x_251, x_2, x_249); -lean_dec_ref(x_246); -x_253 = lean_ctor_get(x_252, 0); -lean_inc(x_253); -x_254 = lean_ctor_get(x_252, 1); -lean_inc(x_254); -lean_dec_ref(x_252); -x_255 = !lean_is_exclusive(x_253); -if (x_255 == 0) -{ -lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_256 = lean_ctor_get(x_253, 0); -x_257 = lean_ctor_get(x_253, 1); -lean_ctor_set(x_1, 1, x_256); -x_258 = l_Lean_IR_reshape(x_257, x_1); -lean_ctor_set(x_253, 1, x_254); -lean_ctor_set(x_253, 0, x_258); -return x_253; -} -else -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_259 = lean_ctor_get(x_253, 0); -x_260 = lean_ctor_get(x_253, 1); -lean_inc(x_260); -lean_inc(x_259); -lean_dec(x_253); -lean_ctor_set(x_1, 1, x_259); -x_261 = l_Lean_IR_reshape(x_260, x_1); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_254); -return x_262; -} -} -else -{ -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_263 = lean_ctor_get(x_1, 0); -x_264 = lean_ctor_get(x_1, 1); -lean_inc(x_264); -lean_inc(x_263); -lean_dec(x_1); -x_265 = l_Lean_IR_ExplicitBoxing_getJPParams(x_263, x_2, x_3); -x_266 = lean_ctor_get(x_265, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_265, 1); -lean_inc(x_267); -lean_dec_ref(x_265); -x_268 = l_Lean_IR_ExplicitBoxing_mkBoxedVersionAux___closed__0; -x_269 = lean_alloc_closure((void*)(l_Lean_IR_ExplicitBoxing_visitFnBody___lam__5___boxed), 3, 2); -lean_closure_set(x_269, 0, x_268); -lean_closure_set(x_269, 1, x_266); -x_270 = l_Lean_IR_ExplicitBoxing_castArgsIfNeededAux(x_264, x_269, x_2, x_267); -lean_dec_ref(x_264); -x_271 = lean_ctor_get(x_270, 0); -lean_inc(x_271); -x_272 = lean_ctor_get(x_270, 1); -lean_inc(x_272); -lean_dec_ref(x_270); -x_273 = lean_ctor_get(x_271, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_271, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - lean_ctor_release(x_271, 1); - x_275 = x_271; +x_250 = lean_ctor_get(x_248, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_248, 1); +lean_inc(x_251); +if (lean_is_exclusive(x_248)) { + lean_ctor_release(x_248, 0); + lean_ctor_release(x_248, 1); + x_252 = x_248; } else { - lean_dec_ref(x_271); - x_275 = lean_box(0); -} -x_276 = lean_alloc_ctor(11, 2, 0); -lean_ctor_set(x_276, 0, x_263); -lean_ctor_set(x_276, 1, x_273); -x_277 = l_Lean_IR_reshape(x_274, x_276); -if (lean_is_scalar(x_275)) { - x_278 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_248); + x_252 = lean_box(0); +} +x_253 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_253, 0, x_240); +lean_ctor_set(x_253, 1, x_250); +x_254 = l_Lean_IR_reshape(x_251, x_253); +if (lean_is_scalar(x_252)) { + x_255 = lean_alloc_ctor(0, 2, 0); } else { - x_278 = x_275; + x_255 = x_252; } -lean_ctor_set(x_278, 0, x_277); -lean_ctor_set(x_278, 1, x_272); -return x_278; +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_249); +return x_255; } } default: { -lean_object* x_279; +lean_object* x_256; lean_dec_ref(x_2); -x_279 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_279, 0, x_1); -lean_ctor_set(x_279, 1, x_3); -return x_279; +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_1); +lean_ctor_set(x_256, 1, x_3); +return x_256; } } } @@ -5099,13 +5022,14 @@ x_8 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR return x_8; } } -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_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_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; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_3); +x_9 = l_Lean_IR_ExplicitBoxing_visitFnBody___lam__0(x_1, x_2, x_8, x_4, x_5, x_6, x_7); +lean_dec_ref(x_6); +return x_9; } } 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) { diff --git a/stage0/stdlib/Lean/Compiler/IR/Checker.c b/stage0/stdlib/Lean/Compiler/IR/Checker.c index e5ee8f69f681..ec24d8265d6b 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Checker.c +++ b/stage0/stdlib/Lean/Compiler/IR/Checker.c @@ -111,6 +111,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_Checker_getMaxCtorFields___boxed(lean_object* static lean_object* l_Lean_IR_Checker_checkPartialApp___closed__0; static lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_IR_Checker_throwCheckerError_spec__0_spec__0___closed__3; static lean_object* l_Lean_IR_Checker_throwCheckerError___redArg___closed__2; +lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign(uint8_t); extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_IR_checkDecls(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_LocalContext_isParam(lean_object*, lean_object*); @@ -3593,447 +3594,448 @@ return x_130; } case 4: { -lean_object* x_131; lean_object* x_132; -x_131 = lean_ctor_get(x_2, 1); -lean_inc(x_131); +uint8_t x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_132 = lean_ctor_get(x_2, 1); +lean_inc(x_132); lean_dec_ref(x_2); -x_132 = l_Lean_IR_Checker_checkObjVar(x_131, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_132) == 0) +x_133 = l_Lean_IR_Checker_checkObjVar(x_132, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_133) == 0) { -uint8_t x_133; -x_133 = !lean_is_exclusive(x_132); -if (x_133 == 0) +uint8_t x_134; +x_134 = !lean_is_exclusive(x_133); +if (x_134 == 0) { -lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_134 = lean_ctor_get(x_132, 0); -lean_dec(x_134); -x_135 = lean_box(5); +lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_135 = lean_ctor_get(x_133, 0); +lean_dec(x_135); +x_136 = l_Lean_IR_IRType_ptrSizedTypeForSign(x_131); lean_inc(x_1); -x_136 = l_Lean_IR_instBEqIRType_beq(x_1, x_135); -if (x_136 == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -lean_free_object(x_132); -x_137 = l_Lean_IR_Checker_checkType___closed__0; -x_138 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_139 = l_Lean_IR_Checker_checkType___closed__1; -x_140 = lean_unsigned_to_nat(0u); -x_141 = l_Std_Format_pretty(x_138, x_139, x_140, x_140); -x_142 = lean_string_append(x_137, x_141); -lean_dec_ref(x_141); -x_143 = l_Lean_IR_Checker_checkVar___closed__2; -x_144 = lean_string_append(x_142, x_143); -x_145 = l_Lean_IR_Checker_throwCheckerError___redArg(x_144, x_3, x_4, x_5, x_6); -lean_dec_ref(x_144); -return x_145; -} -else -{ -lean_object* x_146; +x_137 = l_Lean_IR_instBEqIRType_beq(x_1, x_136); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_free_object(x_133); +x_138 = l_Lean_IR_Checker_checkType___closed__0; +x_139 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_140 = l_Lean_IR_Checker_checkType___closed__1; +x_141 = lean_unsigned_to_nat(0u); +x_142 = l_Std_Format_pretty(x_139, x_140, x_141, x_141); +x_143 = lean_string_append(x_138, x_142); +lean_dec_ref(x_142); +x_144 = l_Lean_IR_Checker_checkVar___closed__2; +x_145 = lean_string_append(x_143, x_144); +x_146 = l_Lean_IR_Checker_throwCheckerError___redArg(x_145, x_3, x_4, x_5, x_6); +lean_dec_ref(x_145); +return x_146; +} +else +{ +lean_object* x_147; lean_dec(x_1); -x_146 = lean_box(0); -lean_ctor_set(x_132, 0, x_146); -return x_132; +x_147 = lean_box(0); +lean_ctor_set(x_133, 0, x_147); +return x_133; } } else { -lean_object* x_147; uint8_t x_148; -lean_dec(x_132); -x_147 = lean_box(5); +lean_object* x_148; uint8_t x_149; +lean_dec(x_133); +x_148 = l_Lean_IR_IRType_ptrSizedTypeForSign(x_131); lean_inc(x_1); -x_148 = l_Lean_IR_instBEqIRType_beq(x_1, x_147); -if (x_148 == 0) -{ -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_149 = l_Lean_IR_Checker_checkType___closed__0; -x_150 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_151 = l_Lean_IR_Checker_checkType___closed__1; -x_152 = lean_unsigned_to_nat(0u); -x_153 = l_Std_Format_pretty(x_150, x_151, x_152, x_152); -x_154 = lean_string_append(x_149, x_153); -lean_dec_ref(x_153); -x_155 = l_Lean_IR_Checker_checkVar___closed__2; -x_156 = lean_string_append(x_154, x_155); -x_157 = l_Lean_IR_Checker_throwCheckerError___redArg(x_156, x_3, x_4, x_5, x_6); -lean_dec_ref(x_156); -return x_157; -} -else -{ -lean_object* x_158; lean_object* x_159; +x_149 = l_Lean_IR_instBEqIRType_beq(x_1, x_148); +if (x_149 == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_150 = l_Lean_IR_Checker_checkType___closed__0; +x_151 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_152 = l_Lean_IR_Checker_checkType___closed__1; +x_153 = lean_unsigned_to_nat(0u); +x_154 = l_Std_Format_pretty(x_151, x_152, x_153, x_153); +x_155 = lean_string_append(x_150, x_154); +lean_dec_ref(x_154); +x_156 = l_Lean_IR_Checker_checkVar___closed__2; +x_157 = lean_string_append(x_155, x_156); +x_158 = l_Lean_IR_Checker_throwCheckerError___redArg(x_157, x_3, x_4, x_5, x_6); +lean_dec_ref(x_157); +return x_158; +} +else +{ +lean_object* x_159; lean_object* x_160; lean_dec(x_1); -x_158 = lean_box(0); -x_159 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_159, 0, x_158); -return x_159; +x_159 = lean_box(0); +x_160 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_160, 0, x_159); +return x_160; } } } else { lean_dec(x_1); -return x_132; +return x_133; } } case 5: { -lean_object* x_160; lean_object* x_161; -x_160 = lean_ctor_get(x_2, 2); -lean_inc(x_160); +lean_object* x_161; lean_object* x_162; +x_161 = lean_ctor_get(x_2, 2); +lean_inc(x_161); lean_dec_ref(x_2); -x_161 = l_Lean_IR_Checker_checkObjVar(x_160, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_161) == 0) +x_162 = l_Lean_IR_Checker_checkObjVar(x_161, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_162) == 0) { -lean_object* x_162; -lean_dec_ref(x_161); -x_162 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); -return x_162; +lean_object* x_163; +lean_dec_ref(x_162); +x_163 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); +return x_163; } else { lean_dec(x_1); -return x_161; +return x_162; } } case 6: { -lean_object* x_163; lean_object* x_164; lean_object* x_165; +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_dec(x_1); -x_163 = lean_ctor_get(x_2, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_164); +x_164 = lean_ctor_get(x_2, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_165); lean_dec_ref(x_2); -x_165 = l_Lean_IR_Checker_checkFullApp(x_163, x_164, x_3, x_4, x_5, x_6); -lean_dec_ref(x_164); -return x_165; +x_166 = l_Lean_IR_Checker_checkFullApp(x_164, x_165, x_3, x_4, x_5, x_6); +lean_dec_ref(x_165); +return x_166; } case 7: { -lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_166 = lean_ctor_get(x_2, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_167); +lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_167 = lean_ctor_get(x_2, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_168); lean_dec_ref(x_2); -x_168 = l_Lean_IR_Checker_checkPartialApp(x_166, x_167, x_3, x_4, x_5, x_6); -lean_dec_ref(x_167); -if (lean_obj_tag(x_168) == 0) -{ -lean_object* x_169; +x_169 = l_Lean_IR_Checker_checkPartialApp(x_167, x_168, x_3, x_4, x_5, x_6); lean_dec_ref(x_168); -x_169 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_169; +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; +lean_dec_ref(x_169); +x_170 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_170; } else { lean_dec(x_1); -return x_168; +return x_169; } } case 8: { -lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_170 = lean_ctor_get(x_2, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_2, 1); -lean_inc_ref(x_171); +lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_171 = lean_ctor_get(x_2, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_2, 1); +lean_inc_ref(x_172); lean_dec_ref(x_2); -x_172 = l_Lean_IR_Checker_checkObjVar(x_170, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_172) == 0) -{ -lean_object* x_173; -lean_dec_ref(x_172); -x_173 = l_Lean_IR_Checker_checkArgs(x_171, x_3, x_4, x_5, x_6); -lean_dec_ref(x_171); +x_173 = l_Lean_IR_Checker_checkObjVar(x_171, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_173) == 0) { lean_object* x_174; lean_dec_ref(x_173); -x_174 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_174; +x_174 = l_Lean_IR_Checker_checkArgs(x_172, x_3, x_4, x_5, x_6); +lean_dec_ref(x_172); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; +lean_dec_ref(x_174); +x_175 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_175; } else { lean_dec(x_1); -return x_173; +return x_174; } } else { -lean_dec_ref(x_171); +lean_dec_ref(x_172); lean_dec(x_1); -return x_172; +return x_173; } } case 9: { -lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_175 = lean_ctor_get(x_2, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_2, 1); +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_2, 0); lean_inc(x_176); +x_177 = lean_ctor_get(x_2, 1); +lean_inc(x_177); lean_dec_ref(x_2); -x_177 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_177) == 0) -{ -lean_object* x_178; -lean_dec_ref(x_177); -lean_inc(x_176); -x_178 = l_Lean_IR_Checker_checkScalarVar(x_176, x_3, x_4, x_5, x_6); +x_178 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_178) == 0) { lean_object* x_179; lean_dec_ref(x_178); -x_179 = l_Lean_IR_Checker_getType(x_176, x_3, x_4, x_5, x_6); +lean_inc(x_177); +x_179 = l_Lean_IR_Checker_checkScalarVar(x_177, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_179) == 0) { -uint8_t x_180; -x_180 = !lean_is_exclusive(x_179); -if (x_180 == 0) -{ -lean_object* x_181; uint8_t x_182; -x_181 = lean_ctor_get(x_179, 0); -lean_inc(x_181); -x_182 = l_Lean_IR_instBEqIRType_beq(x_181, x_175); -if (x_182 == 0) -{ -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -lean_free_object(x_179); -x_183 = l_Lean_IR_Checker_checkType___closed__0; -x_184 = l_Lean_IR_instToFormatIRType___private__1(x_181); -x_185 = l_Lean_IR_Checker_checkType___closed__1; -x_186 = lean_unsigned_to_nat(0u); -x_187 = l_Std_Format_pretty(x_184, x_185, x_186, x_186); -x_188 = lean_string_append(x_183, x_187); -lean_dec_ref(x_187); -x_189 = l_Lean_IR_Checker_checkVar___closed__2; -x_190 = lean_string_append(x_188, x_189); -x_191 = l_Lean_IR_Checker_throwCheckerError___redArg(x_190, x_3, x_4, x_5, x_6); -lean_dec_ref(x_190); -return x_191; -} -else -{ -lean_object* x_192; -lean_dec(x_181); -x_192 = lean_box(0); -lean_ctor_set(x_179, 0, x_192); -return x_179; +lean_object* x_180; +lean_dec_ref(x_179); +x_180 = l_Lean_IR_Checker_getType(x_177, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_180) == 0) +{ +uint8_t x_181; +x_181 = !lean_is_exclusive(x_180); +if (x_181 == 0) +{ +lean_object* x_182; uint8_t x_183; +x_182 = lean_ctor_get(x_180, 0); +lean_inc(x_182); +x_183 = l_Lean_IR_instBEqIRType_beq(x_182, x_176); +if (x_183 == 0) +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_free_object(x_180); +x_184 = l_Lean_IR_Checker_checkType___closed__0; +x_185 = l_Lean_IR_instToFormatIRType___private__1(x_182); +x_186 = l_Lean_IR_Checker_checkType___closed__1; +x_187 = lean_unsigned_to_nat(0u); +x_188 = l_Std_Format_pretty(x_185, x_186, x_187, x_187); +x_189 = lean_string_append(x_184, x_188); +lean_dec_ref(x_188); +x_190 = l_Lean_IR_Checker_checkVar___closed__2; +x_191 = lean_string_append(x_189, x_190); +x_192 = l_Lean_IR_Checker_throwCheckerError___redArg(x_191, x_3, x_4, x_5, x_6); +lean_dec_ref(x_191); +return x_192; +} +else +{ +lean_object* x_193; +lean_dec(x_182); +x_193 = lean_box(0); +lean_ctor_set(x_180, 0, x_193); +return x_180; } } else { -lean_object* x_193; uint8_t x_194; -x_193 = lean_ctor_get(x_179, 0); -lean_inc(x_193); -lean_dec(x_179); -lean_inc(x_193); -x_194 = l_Lean_IR_instBEqIRType_beq(x_193, x_175); -if (x_194 == 0) +lean_object* x_194; uint8_t x_195; +x_194 = lean_ctor_get(x_180, 0); +lean_inc(x_194); +lean_dec(x_180); +lean_inc(x_194); +x_195 = l_Lean_IR_instBEqIRType_beq(x_194, x_176); +if (x_195 == 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; lean_object* x_202; lean_object* x_203; -x_195 = l_Lean_IR_Checker_checkType___closed__0; -x_196 = l_Lean_IR_instToFormatIRType___private__1(x_193); -x_197 = l_Lean_IR_Checker_checkType___closed__1; -x_198 = lean_unsigned_to_nat(0u); -x_199 = l_Std_Format_pretty(x_196, x_197, x_198, x_198); -x_200 = lean_string_append(x_195, x_199); -lean_dec_ref(x_199); -x_201 = l_Lean_IR_Checker_checkVar___closed__2; -x_202 = lean_string_append(x_200, x_201); -x_203 = l_Lean_IR_Checker_throwCheckerError___redArg(x_202, x_3, x_4, x_5, x_6); -lean_dec_ref(x_202); -return x_203; +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_196 = l_Lean_IR_Checker_checkType___closed__0; +x_197 = l_Lean_IR_instToFormatIRType___private__1(x_194); +x_198 = l_Lean_IR_Checker_checkType___closed__1; +x_199 = lean_unsigned_to_nat(0u); +x_200 = l_Std_Format_pretty(x_197, x_198, x_199, x_199); +x_201 = lean_string_append(x_196, x_200); +lean_dec_ref(x_200); +x_202 = l_Lean_IR_Checker_checkVar___closed__2; +x_203 = lean_string_append(x_201, x_202); +x_204 = l_Lean_IR_Checker_throwCheckerError___redArg(x_203, x_3, x_4, x_5, x_6); +lean_dec_ref(x_203); +return x_204; } else { -lean_object* x_204; lean_object* x_205; -lean_dec(x_193); -x_204 = lean_box(0); -x_205 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_205, 0, x_204); -return x_205; +lean_object* x_205; lean_object* x_206; +lean_dec(x_194); +x_205 = lean_box(0); +x_206 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_206, 0, x_205); +return x_206; } } } else { -uint8_t x_206; -lean_dec(x_175); -x_206 = !lean_is_exclusive(x_179); -if (x_206 == 0) +uint8_t x_207; +lean_dec(x_176); +x_207 = !lean_is_exclusive(x_180); +if (x_207 == 0) { -return x_179; +return x_180; } else { -lean_object* x_207; lean_object* x_208; -x_207 = lean_ctor_get(x_179, 0); -lean_inc(x_207); -lean_dec(x_179); -x_208 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_208, 0, x_207); -return x_208; +lean_object* x_208; lean_object* x_209; +x_208 = lean_ctor_get(x_180, 0); +lean_inc(x_208); +lean_dec(x_180); +x_209 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_209, 0, x_208); +return x_209; } } } else { +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -return x_178; +return x_179; } } else { +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -return x_177; +return x_178; } } case 10: { -lean_object* x_209; lean_object* x_210; -x_209 = lean_ctor_get(x_2, 0); -lean_inc(x_209); +lean_object* x_210; lean_object* x_211; +x_210 = lean_ctor_get(x_2, 0); +lean_inc(x_210); lean_dec_ref(x_2); -x_210 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_210) == 0) +x_211 = l_Lean_IR_Checker_checkScalarType(x_1, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_211) == 0) { -lean_object* x_211; -lean_dec_ref(x_210); -x_211 = l_Lean_IR_Checker_checkObjVar(x_209, x_3, x_4, x_5, x_6); -return x_211; +lean_object* x_212; +lean_dec_ref(x_211); +x_212 = l_Lean_IR_Checker_checkObjVar(x_210, x_3, x_4, x_5, x_6); +return x_212; } else { -lean_dec(x_209); -return x_210; +lean_dec(x_210); +return x_211; } } case 11: { -lean_object* x_212; -x_212 = lean_ctor_get(x_2, 0); -lean_inc_ref(x_212); +lean_object* x_213; +x_213 = lean_ctor_get(x_2, 0); +lean_inc_ref(x_213); lean_dec_ref(x_2); -if (lean_obj_tag(x_212) == 0) +if (lean_obj_tag(x_213) == 0) { -uint8_t x_213; +uint8_t x_214; lean_dec(x_1); -x_213 = !lean_is_exclusive(x_212); -if (x_213 == 0) -{ -lean_object* x_214; lean_object* x_215; -x_214 = lean_ctor_get(x_212, 0); -lean_dec(x_214); -x_215 = lean_box(0); -lean_ctor_set(x_212, 0, x_215); -return x_212; +x_214 = !lean_is_exclusive(x_213); +if (x_214 == 0) +{ +lean_object* x_215; lean_object* x_216; +x_215 = lean_ctor_get(x_213, 0); +lean_dec(x_215); +x_216 = lean_box(0); +lean_ctor_set(x_213, 0, x_216); +return x_213; } else { -lean_object* x_216; lean_object* x_217; -lean_dec(x_212); -x_216 = lean_box(0); -x_217 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_217, 0, x_216); -return x_217; +lean_object* x_217; lean_object* x_218; +lean_dec(x_213); +x_217 = lean_box(0); +x_218 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_218, 0, x_217); +return x_218; } } else { -lean_object* x_218; -lean_dec_ref(x_212); -x_218 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); -return x_218; +lean_object* x_219; +lean_dec_ref(x_213); +x_219 = l_Lean_IR_Checker_checkObjType(x_1, x_3, x_4, x_5, x_6); +return x_219; } } default: { -lean_object* x_219; lean_object* x_220; -x_219 = lean_ctor_get(x_2, 0); -lean_inc(x_219); +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_2, 0); +lean_inc(x_220); lean_dec_ref(x_2); -x_220 = l_Lean_IR_Checker_checkObjVar(x_219, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_220) == 0) +x_221 = l_Lean_IR_Checker_checkObjVar(x_220, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_221) == 0) { -uint8_t x_221; -x_221 = !lean_is_exclusive(x_220); -if (x_221 == 0) +uint8_t x_222; +x_222 = !lean_is_exclusive(x_221); +if (x_222 == 0) { -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_object* x_223; lean_object* x_224; uint8_t x_225; +x_223 = lean_ctor_get(x_221, 0); +lean_dec(x_223); +x_224 = lean_box(1); lean_inc(x_1); -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_checkType___closed__0; -x_226 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_227 = l_Lean_IR_Checker_checkType___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); -lean_dec_ref(x_232); -return x_233; -} -else -{ -lean_object* x_234; +x_225 = l_Lean_IR_instBEqIRType_beq(x_1, 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; lean_object* x_234; +lean_free_object(x_221); +x_226 = l_Lean_IR_Checker_checkType___closed__0; +x_227 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_228 = l_Lean_IR_Checker_checkType___closed__1; +x_229 = lean_unsigned_to_nat(0u); +x_230 = l_Std_Format_pretty(x_227, x_228, x_229, x_229); +x_231 = lean_string_append(x_226, x_230); +lean_dec_ref(x_230); +x_232 = l_Lean_IR_Checker_checkVar___closed__2; +x_233 = lean_string_append(x_231, x_232); +x_234 = l_Lean_IR_Checker_throwCheckerError___redArg(x_233, x_3, x_4, x_5, x_6); +lean_dec_ref(x_233); +return x_234; +} +else +{ +lean_object* x_235; lean_dec(x_1); -x_234 = lean_box(0); -lean_ctor_set(x_220, 0, x_234); -return x_220; +x_235 = lean_box(0); +lean_ctor_set(x_221, 0, x_235); +return x_221; } } else { -lean_object* x_235; uint8_t x_236; -lean_dec(x_220); -x_235 = lean_box(1); +lean_object* x_236; uint8_t x_237; +lean_dec(x_221); +x_236 = lean_box(1); lean_inc(x_1); -x_236 = l_Lean_IR_instBEqIRType_beq(x_1, x_235); -if (x_236 == 0) -{ -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_237 = l_Lean_IR_Checker_checkType___closed__0; -x_238 = l_Lean_IR_instToFormatIRType___private__1(x_1); -x_239 = l_Lean_IR_Checker_checkType___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); -lean_dec_ref(x_244); -return x_245; -} -else -{ -lean_object* x_246; lean_object* x_247; +x_237 = l_Lean_IR_instBEqIRType_beq(x_1, 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; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_238 = l_Lean_IR_Checker_checkType___closed__0; +x_239 = l_Lean_IR_instToFormatIRType___private__1(x_1); +x_240 = l_Lean_IR_Checker_checkType___closed__1; +x_241 = lean_unsigned_to_nat(0u); +x_242 = l_Std_Format_pretty(x_239, x_240, x_241, x_241); +x_243 = lean_string_append(x_238, x_242); +lean_dec_ref(x_242); +x_244 = l_Lean_IR_Checker_checkVar___closed__2; +x_245 = lean_string_append(x_243, x_244); +x_246 = l_Lean_IR_Checker_throwCheckerError___redArg(x_245, x_3, x_4, x_5, x_6); +lean_dec_ref(x_245); +return x_246; +} +else +{ +lean_object* x_247; lean_object* x_248; lean_dec(x_1); -x_246 = lean_box(0); -x_247 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_247, 0, x_246); -return x_247; +x_247 = lean_box(0); +x_248 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_248, 0, x_247); +return x_248; } } } else { lean_dec(x_1); -return x_220; +return x_221; } } } diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index 0409dcc91625..12b70a0e5d1b 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -55,6 +55,7 @@ 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*); uint8_t l_Lean_IR_instBEqJoinPointId_beq(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +static lean_object* l_Lean_IR_EmitC_emitUProj___redArg___closed__1; 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___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*); @@ -67,7 +68,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitArg___boxed(lean_object*, lean_obje static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__5; 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_Lean_IR_EmitC_emitUProj(lean_object*, lean_object*, uint8_t, 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*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__4; static lean_object* l_Lean_IR_EmitC_emitLn___redArg___closed__0; @@ -136,6 +137,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSetTag___redArg(lean_object*, lean_ 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_emitUSet___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_emitFnDeclAux(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); @@ -170,6 +172,7 @@ 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_toCType___closed__13; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__0; lean_object* lean_nat_shiftr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitInitFn_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -191,7 +194,7 @@ lean_object* l_ReaderT_instMonad___redArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_isIf___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__0(lean_object*, size_t, size_t, lean_object*, lean_object*, 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_emitUSet___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__9; 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; @@ -226,6 +229,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Fold_0__Nat_anyTR_loop___at__ 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_Lean_IR_EmitC_emitLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_toCType___closed__14; static lean_object* l_Lean_IR_EmitC_emitDec___redArg___closed__1; 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__1_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__18; @@ -237,6 +241,7 @@ 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_emitSProj___redArg___closed__6; static lean_object* l_Lean_IR_EmitC_leanMainFn___closed__0; uint8_t l_Lean_IR_usesModuleFrom(lean_object*, lean_object*); extern lean_object* l_Lean_IR_instInhabitedParam_default; @@ -250,6 +255,7 @@ uint32_t l_Nat_digitChar(lean_object*); static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCName___boxed(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +static lean_object* l_Lean_IR_EmitC_emitUSet___redArg___closed__1; 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; @@ -260,10 +266,11 @@ uint8_t l_Lean_isIOUnitBuiltinInitFn(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__9; static lean_object* l_Lean_IR_EmitC_emitDeclAux___lam__0___closed__0; +static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__7; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__3; static 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__1_spec__1___redArg___closed__0; 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*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitIsShared___redArg___closed__0; LEAN_EXPORT uint8_t l_Lean_IR_EmitC_overwriteParam(lean_object*, lean_object*); static lean_object* l_String_foldlAux___at___00Lean_IR_EmitC_quoteString_spec__0___closed__2; @@ -291,7 +298,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileFooter(lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFileHeader___boxed(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_hasMainFn___boxed(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_emitUSet___boxed(lean_object*, 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; @@ -303,6 +310,7 @@ static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLn(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_Lean_IR_EmitC_emitTag___redArg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__7; uint8_t l_Lean_IR_IRType_isObj(lean_object*); static lean_object* l_Lean_IR_EmitC_emitNumLit___redArg___closed__2; static lean_object* l_Lean_IR_EmitC_emitCase___closed__1; @@ -340,6 +348,7 @@ static lean_object* l_Lean_IR_EmitC_emitCtor___closed__0; static lean_object* l_Lean_IR_EmitC_emitJmp___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_throwInvalidExportName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitMarkPersistent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__8; LEAN_EXPORT uint8_t l_Lean_IR_EmitC_hasMainFn___lam__0(lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitCase_spec__0___closed__1; @@ -349,7 +358,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLhs___redArg(lean_object*, lean_obj 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*); 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_Lean_IR_EmitC_emitUProj___redArg(lean_object*, lean_object*, uint8_t, 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*); static lean_object* l_Lean_IR_EmitC_emitJmp___closed__1; static lean_object* l_Lean_IR_EmitC_toCInitName___closed__0; @@ -358,6 +367,7 @@ static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__36; 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; +static lean_object* l_Lean_IR_EmitC_toCType___closed__17; 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*); @@ -387,6 +397,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_foldM_loop___a 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*); 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*); +static lean_object* l_Lean_IR_EmitC_toCType___closed__15; 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; @@ -422,6 +433,7 @@ 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_Lean_IR_EmitC_emitJmp(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*); +static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__6; 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_Lean_IR_EmitC_emitFnDeclAux___lam__0___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitFileHeader___closed__2; @@ -460,6 +472,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitPartialApp___boxed(lean_object*, le LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitBlock___boxed(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; +static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8; static lean_object* l_Lean_IR_EmitC_emitDeclAux___lam__0___closed__1; static lean_object* l_String_foldlAux___at___00Lean_IR_EmitC_quoteString_spec__0___closed__3; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__7; @@ -467,12 +480,15 @@ 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; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtorScalarSize(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*); +static lean_object* l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDec___redArg___boxed(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*); static lean_object* l_Lean_IR_EmitC_emitSProj___redArg___closed__2; @@ -507,6 +523,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe static lean_object* l_Lean_IR_EmitC_toCType___closed__6; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__13; uint8_t l_Lean_IR_isTailCallTo(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__10; 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_Lean_IR_EmitC_emitDel(lean_object*, lean_object*, lean_object*); @@ -562,7 +579,8 @@ lean_object* l_Nat_nextPowerOfTwo(lean_object*); 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*); +static lean_object* l_Lean_IR_EmitC_emitSSet___redArg___closed__9; +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___boxed(lean_object*, lean_object*, lean_object*, 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_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitC_getJPParams_spec__0(lean_object*, lean_object*, lean_object*); @@ -636,6 +654,7 @@ static lean_object* l_Lean_IR_EmitC_emitUProj___redArg___closed__0; static lean_object* l_Lean_IR_EmitC_emitLns___redArg___closed__5; 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_emitSProj___redArg___closed__8; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitVDecl___boxed(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__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -660,6 +679,7 @@ 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*); +static lean_object* l_Lean_IR_EmitC_toCType___closed__16; static lean_object* l_Lean_IR_EmitC_overwriteParam___closed__0; static lean_object* l_Lean_IR_EmitC_emitMainFn___closed__43; 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*); @@ -1308,7 +1328,7 @@ static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__10() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_EmitC_toCType___closed__9; x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(69u); +x_3 = lean_unsigned_to_nat(74u); 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); @@ -1321,7 +1341,7 @@ static lean_object* _init_l_Lean_IR_EmitC_toCType___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_EmitC_toCType___closed__9; x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(70u); +x_3 = lean_unsigned_to_nat(75u); 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); @@ -1332,6 +1352,46 @@ static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__12() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("int8_t", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("int16_t", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("int32_t", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("int64_t", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ptrdiff_t", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_toCType___closed__17() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("lean_object*", 12, 12); return x_1; } @@ -1396,12 +1456,42 @@ 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; } -default: +case 14: { lean_object* x_13; x_13 = l_Lean_IR_EmitC_toCType___closed__12; return x_13; } +case 15: +{ +lean_object* x_14; +x_14 = l_Lean_IR_EmitC_toCType___closed__13; +return x_14; +} +case 16: +{ +lean_object* x_15; +x_15 = l_Lean_IR_EmitC_toCType___closed__14; +return x_15; +} +case 17: +{ +lean_object* x_16; +x_16 = l_Lean_IR_EmitC_toCType___closed__15; +return x_16; +} +case 18: +{ +lean_object* x_17; +x_17 = l_Lean_IR_EmitC_toCType___closed__16; +return x_17; +} +default: +{ +lean_object* x_18; +x_18 = l_Lean_IR_EmitC_toCType___closed__17; +return x_18; +} } } } @@ -2268,10 +2358,10 @@ return x_14; if (x_21 == 0) { lean_object* x_22; lean_object* x_23; -lean_inc(x_20); -x_22 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(x_18, x_16, x_20, x_20, x_17); -lean_dec(x_20); -lean_dec_ref(x_18); +lean_inc(x_17); +x_22 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitFnDeclAux_spec__0___redArg(x_20, x_18, x_17, x_17, x_16); +lean_dec(x_17); +lean_dec_ref(x_20); x_23 = lean_ctor_get(x_22, 1); lean_inc(x_23); lean_dec_ref(x_22); @@ -2282,10 +2372,10 @@ goto block_15; else { lean_object* x_24; lean_object* x_25; -lean_dec(x_20); -lean_dec_ref(x_18); +lean_dec_ref(x_20); +lean_dec(x_17); x_24 = l_Lean_IR_EmitC_emitFnDeclAux___closed__1; -x_25 = lean_string_append(x_17, x_24); +x_25 = lean_string_append(x_16, x_24); x_9 = x_19; x_10 = x_25; goto block_15; @@ -2299,25 +2389,25 @@ x_33 = lean_array_get_size(x_31); x_34 = lean_nat_dec_lt(x_32, x_33); if (x_34 == 0) { -lean_dec(x_29); +lean_dec(x_28); x_16 = x_27; -x_17 = x_28; -x_18 = x_31; +x_17 = x_33; +x_18 = x_29; x_19 = x_30; -x_20 = x_33; +x_20 = x_31; x_21 = x_34; goto block_26; } else { uint8_t x_35; -x_35 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_29); -lean_dec(x_29); +x_35 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_28); +lean_dec(x_28); x_16 = x_27; -x_17 = x_28; -x_18 = x_31; +x_17 = x_33; +x_18 = x_29; x_19 = x_30; -x_20 = x_33; +x_20 = x_31; x_21 = x_35; goto block_26; } @@ -2331,8 +2421,8 @@ x_42 = l_Lean_isExternC(x_7, x_41); if (x_42 == 0) { x_27 = x_37; -x_28 = x_38; -x_29 = x_41; +x_28 = x_41; +x_29 = x_38; x_30 = x_39; x_31 = x_40; goto block_36; @@ -2341,15 +2431,15 @@ else { 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_37); -x_45 = lean_nat_dec_lt(x_37, x_43); +x_44 = lean_mk_empty_array_with_capacity(x_38); +x_45 = lean_nat_dec_lt(x_38, x_43); if (x_45 == 0) { lean_dec(x_43); lean_dec_ref(x_40); x_27 = x_37; -x_28 = x_38; -x_29 = x_41; +x_28 = x_41; +x_29 = x_38; x_30 = x_39; x_31 = x_44; goto block_36; @@ -2363,8 +2453,8 @@ if (x_46 == 0) lean_dec(x_43); lean_dec_ref(x_40); x_27 = x_37; -x_28 = x_38; -x_29 = x_41; +x_28 = x_41; +x_29 = x_38; x_30 = x_39; x_31 = x_44; goto block_36; @@ -2378,8 +2468,8 @@ lean_dec(x_43); x_49 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__1(x_40, x_47, x_48, x_44); lean_dec_ref(x_40); x_27 = x_37; -x_28 = x_38; -x_29 = x_41; +x_28 = x_41; +x_29 = x_38; x_30 = x_39; x_31 = x_49; goto block_36; @@ -2412,8 +2502,8 @@ if (x_66 == 0) { lean_dec(x_64); lean_dec_ref(x_51); -x_37 = x_63; -x_38 = x_62; +x_37 = x_62; +x_38 = x_63; x_39 = x_52; x_40 = x_65; goto block_50; @@ -2426,8 +2516,8 @@ if (x_67 == 0) { lean_dec(x_64); lean_dec_ref(x_51); -x_37 = x_63; -x_38 = x_62; +x_37 = x_62; +x_38 = x_63; x_39 = x_52; x_40 = x_65; goto block_50; @@ -2440,8 +2530,8 @@ x_69 = lean_usize_of_nat(x_64); lean_dec(x_64); x_70 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_51, x_68, x_69, x_65); lean_dec_ref(x_51); -x_37 = x_63; -x_38 = x_62; +x_37 = x_62; +x_38 = x_63; x_39 = x_52; x_40 = x_70; goto block_50; @@ -3593,9 +3683,9 @@ x_108 = l_Lean_IR_EmitC_emitMainFn___closed__29; x_109 = lean_string_append(x_104, x_108); x_110 = lean_string_append(x_109, x_88); x_60 = x_101; -x_61 = x_80; +x_61 = x_88; x_62 = x_100; -x_63 = x_88; +x_63 = x_80; x_64 = x_81; x_65 = x_110; goto block_74; @@ -3612,9 +3702,9 @@ x_114 = l_Lean_IR_EmitC_emitMainFn___closed__45; x_115 = lean_string_append(x_113, x_114); x_116 = lean_string_append(x_115, x_88); x_60 = x_101; -x_61 = x_80; +x_61 = x_88; x_62 = x_100; -x_63 = x_88; +x_63 = x_80; x_64 = x_81; x_65 = x_116; goto block_74; @@ -3774,10 +3864,10 @@ 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_4); +lean_inc_ref(x_9); x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_4); -lean_ctor_set(x_18, 1, x_3); +lean_ctor_set(x_18, 0, x_9); +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); @@ -3800,12 +3890,12 @@ 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, 0, x_6); lean_ctor_set(x_26, 1, x_25); x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_7); +lean_ctor_set(x_27, 0, x_10); 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_10); +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) @@ -3814,11 +3904,11 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean 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_4); -lean_dec_ref(x_4); +x_32 = lean_string_append(x_30, x_9); +lean_dec_ref(x_9); x_33 = lean_box(0); -x_34 = lean_string_append(x_32, x_6); -lean_dec_ref(x_6); +x_34 = lean_string_append(x_32, x_4); +lean_dec_ref(x_4); lean_ctor_set(x_28, 1, x_34); lean_ctor_set(x_28, 0, x_33); return x_28; @@ -3829,11 +3919,11 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean x_35 = lean_ctor_get(x_28, 1); lean_inc(x_35); lean_dec(x_28); -x_36 = lean_string_append(x_35, x_4); -lean_dec_ref(x_4); +x_36 = lean_string_append(x_35, x_9); +lean_dec_ref(x_9); x_37 = lean_box(0); -x_38 = lean_string_append(x_36, x_6); -lean_dec_ref(x_6); +x_38 = lean_string_append(x_36, x_4); +lean_dec_ref(x_4); x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); @@ -3861,13 +3951,13 @@ 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_4 = x_43; +x_5 = x_42; x_6 = x_44; -x_7 = x_51; +x_7 = x_45; x_8 = x_52; -x_9 = x_45; -x_10 = x_46; +x_9 = x_46; +x_10 = x_51; x_11 = x_56; goto block_40; } @@ -3876,13 +3966,13 @@ 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_4 = x_43; +x_5 = x_42; x_6 = x_44; -x_7 = x_51; +x_7 = x_45; x_8 = x_52; -x_9 = x_45; -x_10 = x_46; +x_9 = x_46; +x_10 = x_51; x_11 = x_57; goto block_40; } @@ -3892,20 +3982,20 @@ goto block_40; 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_68 = lean_string_append(x_67, x_61); x_69 = 0; -x_70 = l_Lean_Environment_find_x3f(x_61, x_59, x_69); +x_70 = l_Lean_Environment_find_x3f(x_63, 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__16; x_72 = l_panic___at___00Lean_IR_EmitC_emitMainFn_spec__3(x_71); -x_41 = x_60; -x_42 = x_66; -x_43 = x_62; -x_44 = x_63; -x_45 = x_64; -x_46 = x_68; +x_41 = x_64; +x_42 = x_60; +x_43 = x_61; +x_44 = x_62; +x_45 = x_68; +x_46 = x_66; x_47 = x_72; goto block_58; } @@ -3915,12 +4005,12 @@ 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_66; -x_43 = x_62; -x_44 = x_63; -x_45 = x_64; -x_46 = x_68; +x_41 = x_64; +x_42 = x_60; +x_43 = x_61; +x_44 = x_62; +x_45 = x_68; +x_46 = x_66; x_47 = x_73; goto block_58; } @@ -5865,12 +5955,37 @@ x_1 = lean_mk_string_unchecked("lean_ctor_set_usize(", 20, 20); return x_1; } } -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) { +static lean_object* _init_l_Lean_IR_EmitC_emitUSet___redArg___closed__1() { _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; lean_object* x_17; lean_object* x_18; 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); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_isize(", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet___redArg(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; +if (x_3 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_IR_EmitC_emitUSet___redArg___closed__0; +x_27 = lean_string_append(x_5, x_26); +x_6 = x_27; +goto block_25; +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = l_Lean_IR_EmitC_emitUSet___redArg___closed__1; +x_29 = lean_string_append(x_5, x_28); +x_6 = x_29; +goto block_25; +} +block_25: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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_7 = l_Lean_IR_EmitC_argToCString___closed__0; x_8 = l_Nat_reprFast(x_1); x_9 = lean_string_append(x_7, x_8); @@ -5883,7 +5998,7 @@ 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_16 = l_Nat_reprFast(x_4); x_17 = lean_string_append(x_7, x_16); lean_dec_ref(x_16); x_18 = lean_string_append(x_15, x_17); @@ -5899,21 +6014,32 @@ lean_ctor_set(x_24, 1, x_23); return x_24; } } -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_EXPORT lean_object* l_Lean_IR_EmitC_emitUSet(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUSet___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_7; +x_7 = l_Lean_IR_EmitC_emitUSet___redArg(x_1, x_2, x_3, x_4, x_6); +return x_7; } } -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_Lean_IR_EmitC_emitUSet___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_IR_EmitC_emitUSet(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitUSet___redArg(x_1, x_2, x_6, x_4, x_5); +return x_7; +} +} +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_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +x_8 = l_Lean_IR_EmitC_emitUSet(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec_ref(x_5); +return x_8; } } static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__0() { @@ -5968,6 +6094,38 @@ static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__6() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int8", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int16", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int32", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int64", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__10() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("invalid instruction", 19, 19); return x_1; } @@ -6025,18 +6183,50 @@ x_52 = lean_string_append(x_6, x_51); x_7 = x_52; goto block_40; } -default: +case 14: { lean_object* x_53; lean_object* x_54; +x_53 = l_Lean_IR_EmitC_emitSSet___redArg___closed__6; +x_54 = lean_string_append(x_6, x_53); +x_7 = x_54; +goto block_40; +} +case 15: +{ +lean_object* x_55; lean_object* x_56; +x_55 = l_Lean_IR_EmitC_emitSSet___redArg___closed__7; +x_56 = lean_string_append(x_6, x_55); +x_7 = x_56; +goto block_40; +} +case 16: +{ +lean_object* x_57; lean_object* x_58; +x_57 = l_Lean_IR_EmitC_emitSSet___redArg___closed__8; +x_58 = lean_string_append(x_6, x_57); +x_7 = x_58; +goto block_40; +} +case 17: +{ +lean_object* x_59; lean_object* x_60; +x_59 = l_Lean_IR_EmitC_emitSSet___redArg___closed__9; +x_60 = lean_string_append(x_6, x_59); +x_7 = x_60; +goto block_40; +} +default: +{ +lean_object* x_61; lean_object* x_62; 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; +x_61 = l_Lean_IR_EmitC_emitSSet___redArg___closed__10; +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_6); +return x_62; } } block_40: @@ -7378,86 +7568,92 @@ x_1 = lean_mk_string_unchecked("lean_ctor_get_usize(", 20, 20); return x_1; } } -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) { +static lean_object* _init_l_Lean_IR_EmitC_emitUProj___redArg___closed__1() { _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_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_isize(", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___redArg(lean_object* x_1, lean_object* x_2, uint8_t 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; lean_object* x_16; lean_object* x_17; 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___redArg___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; +lean_object* x_6; lean_object* x_22; +x_22 = l_Lean_IR_EmitC_emitLhs___redArg(x_1, x_5); +if (x_3 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec_ref(x_22); +x_24 = l_Lean_IR_EmitC_emitUProj___redArg___closed__0; +x_25 = lean_string_append(x_23, x_24); +x_6 = x_25; +goto block_21; } 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___redArg___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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec_ref(x_22); +x_27 = l_Lean_IR_EmitC_emitUProj___redArg___closed__1; +x_28 = lean_string_append(x_26, x_27); +x_6 = x_28; +goto block_21; +} +block_21: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; 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_7 = l_Lean_IR_EmitC_argToCString___closed__0; +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_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 = l_Lean_IR_EmitC_emitInc___redArg___closed__0; +x_16 = lean_string_append(x_14, x_15); +x_17 = l_Lean_IR_EmitC_emitLn___redArg___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_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 lean_object* l_Lean_IR_EmitC_emitUProj(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; -x_6 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_2, x_3, x_5); -return x_6; +lean_object* x_7; +x_7 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_2, x_3, x_4, x_6); +return x_7; } } -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) { +LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitUProj___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_IR_EmitC_emitUProj(x_1, x_2, x_3, x_4, x_5); -lean_dec_ref(x_4); -return x_6; +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_3); +x_7 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_2, x_6, x_4, x_5); +return x_7; +} +} +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, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +x_8 = l_Lean_IR_EmitC_emitUProj(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec_ref(x_5); +return x_8; } } static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__0() { @@ -7508,6 +7704,38 @@ x_1 = lean_mk_string_unchecked("lean_ctor_get_float32", 21, 21); return x_1; } } +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int8", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int16", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int32", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int64", 19, 19); +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) { _start: { @@ -7580,34 +7808,78 @@ x_51 = lean_string_append(x_49, x_50); x_7 = x_51; goto block_32; } +case 14: +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_33, 1); +lean_inc(x_52); +lean_dec_ref(x_33); +x_53 = l_Lean_IR_EmitC_emitSProj___redArg___closed__6; +x_54 = lean_string_append(x_52, x_53); +x_7 = x_54; +goto block_32; +} +case 15: +{ +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_ref(x_33); +x_56 = l_Lean_IR_EmitC_emitSProj___redArg___closed__7; +x_57 = lean_string_append(x_55, x_56); +x_7 = x_57; +goto block_32; +} +case 16: +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_33, 1); +lean_inc(x_58); +lean_dec_ref(x_33); +x_59 = l_Lean_IR_EmitC_emitSProj___redArg___closed__8; +x_60 = lean_string_append(x_58, x_59); +x_7 = x_60; +goto block_32; +} +case 17: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_33, 1); +lean_inc(x_61); +lean_dec_ref(x_33); +x_62 = l_Lean_IR_EmitC_emitSProj___redArg___closed__9; +x_63 = lean_string_append(x_61, x_62); +x_7 = x_63; +goto block_32; +} default: { -uint8_t x_52; +uint8_t x_64; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_52 = !lean_is_exclusive(x_33); -if (x_52 == 0) +x_64 = !lean_is_exclusive(x_33); +if (x_64 == 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_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_33, 0); +lean_dec(x_65); +x_66 = l_Lean_IR_EmitC_emitSSet___redArg___closed__10; lean_ctor_set_tag(x_33, 1); -lean_ctor_set(x_33, 0, x_54); +lean_ctor_set(x_33, 0, x_66); return x_33; } else { -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_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_33, 1); +lean_inc(x_67); 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; +x_68 = l_Lean_IR_EmitC_emitSSet___redArg___closed__10; +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_67); +return x_69; } } } @@ -8884,6 +9156,30 @@ static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__5() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_int32", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_int64", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_isize", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("lean_box", 8, 8); return x_1; } @@ -8947,7 +9243,7 @@ lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); return x_22; } -default: +case 16: { 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; @@ -8958,6 +9254,39 @@ lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); return x_26; } +case 17: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__6; +x_28 = lean_box(0); +x_29 = lean_string_append(x_2, x_27); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +case 18: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7; +x_32 = lean_box(0); +x_33 = lean_string_append(x_2, 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; +} +default: +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8; +x_36 = lean_box(0); +x_37 = lean_string_append(x_2, 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; +} } } } @@ -9889,109 +10218,110 @@ return x_19; } case 4: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; 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); +x_21 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_22 = lean_ctor_get(x_3, 1); +lean_inc(x_22); lean_dec_ref(x_3); -x_22 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_20, x_21, x_5); -return x_22; +x_23 = l_Lean_IR_EmitC_emitUProj___redArg(x_1, x_20, x_21, x_22, x_5); +return x_23; } case 5: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_3, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_3, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_3, 0); lean_inc(x_24); -x_25 = lean_ctor_get(x_3, 2); +x_25 = lean_ctor_get(x_3, 1); lean_inc(x_25); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); 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); +x_27 = l_Lean_IR_EmitC_emitSProj___redArg(x_1, x_2, x_24, x_25, x_26, x_5); lean_dec(x_2); -return x_26; +return x_27; } case 6: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_28; lean_object* x_29; lean_object* x_30; 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); +x_28 = lean_ctor_get(x_3, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_29); 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; +x_30 = l_Lean_IR_EmitC_emitFullApp(x_1, x_28, x_29, x_4, x_5); +return x_30; } case 7: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_31; lean_object* x_32; lean_object* x_33; 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); +x_31 = lean_ctor_get(x_3, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_32); 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; +x_33 = l_Lean_IR_EmitC_emitPartialApp(x_1, x_31, x_32, x_4, x_5); +lean_dec_ref(x_32); +return x_33; } case 8: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_34; lean_object* x_35; lean_object* x_36; 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); +x_34 = lean_ctor_get(x_3, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_35); 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_34); -return x_35; +x_36 = l_Lean_IR_EmitC_emitApp(x_1, x_34, x_35, x_4, x_5); +lean_dec_ref(x_35); +return x_36; } case 9: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_dec(x_2); -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, 0); lean_inc(x_37); +x_38 = lean_ctor_get(x_3, 1); +lean_inc(x_38); 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; +x_39 = l_Lean_IR_EmitC_emitBox___redArg(x_1, x_38, x_37, x_5); +lean_dec(x_37); +return x_39; } case 10: { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_3, 0); -lean_inc(x_39); +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_3, 0); +lean_inc(x_40); lean_dec_ref(x_3); -x_40 = l_Lean_IR_EmitC_emitUnbox___redArg(x_1, x_2, x_39, x_5); +x_41 = l_Lean_IR_EmitC_emitUnbox___redArg(x_1, x_2, x_40, x_5); lean_dec(x_2); -return x_40; +return x_41; } case 11: { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_3, 0); -lean_inc_ref(x_41); +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_3, 0); +lean_inc_ref(x_42); lean_dec_ref(x_3); -x_42 = l_Lean_IR_EmitC_emitLit___redArg(x_1, x_2, x_41, x_5); -return x_42; +x_43 = l_Lean_IR_EmitC_emitLit___redArg(x_1, x_2, x_42, x_5); +return x_43; } default: { -lean_object* x_43; lean_object* x_44; +lean_object* x_44; lean_object* x_45; lean_dec(x_2); -x_43 = lean_ctor_get(x_3, 0); -lean_inc(x_43); +x_44 = lean_ctor_get(x_3, 0); +lean_inc(x_44); lean_dec_ref(x_3); -x_44 = l_Lean_IR_EmitC_emitIsShared___redArg(x_1, x_43, x_5); -return x_44; +x_45 = l_Lean_IR_EmitC_emitIsShared___redArg(x_1, x_44, x_5); +return x_45; } } } @@ -11385,220 +11715,221 @@ goto _start; } 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; +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; x_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_32 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_33 = lean_ctor_get(x_1, 3); +lean_inc(x_33); lean_dec_ref(x_1); -x_33 = l_Lean_IR_EmitC_emitUSet___redArg(x_29, x_30, x_31, x_3); -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; +x_34 = l_Lean_IR_EmitC_emitUSet___redArg(x_29, x_30, x_32, x_31, x_3); +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; } case 5: { -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_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, 0); lean_inc(x_37); -x_38 = lean_ctor_get(x_1, 2); +x_38 = lean_ctor_get(x_1, 1); lean_inc(x_38); -x_39 = lean_ctor_get(x_1, 3); +x_39 = lean_ctor_get(x_1, 2); lean_inc(x_39); -x_40 = lean_ctor_get(x_1, 4); +x_40 = lean_ctor_get(x_1, 3); lean_inc(x_40); -x_41 = lean_ctor_get(x_1, 5); +x_41 = lean_ctor_get(x_1, 4); lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 5); +lean_inc(x_42); lean_dec_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) +x_43 = l_Lean_IR_EmitC_emitSSet___redArg(x_37, x_38, x_39, x_40, x_41, x_3); +lean_dec(x_41); +if (lean_obj_tag(x_43) == 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; +lean_object* x_44; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec_ref(x_43); +x_1 = x_42; +x_3 = x_44; goto _start; } else { -lean_dec(x_41); -return x_42; +lean_dec(x_42); +return x_43; } } case 6: { -uint8_t x_45; -x_45 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -if (x_45 == 0) +uint8_t x_46; +x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_46 == 0) { -lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = lean_ctor_get(x_1, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_1, 1); +lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_47 = lean_ctor_get(x_1, 0); 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); +x_48 = lean_ctor_get(x_1, 1); +lean_inc(x_48); +x_49 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_50 = lean_ctor_get(x_1, 2); +lean_inc(x_50); lean_dec_ref(x_1); -x_50 = l_Lean_IR_EmitC_emitInc___redArg(x_46, x_47, x_48, x_3); -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; +x_51 = l_Lean_IR_EmitC_emitInc___redArg(x_47, x_48, x_49, x_3); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec_ref(x_51); +x_1 = x_50; +x_3 = x_52; goto _start; } else { -lean_object* x_53; -x_53 = lean_ctor_get(x_1, 2); -lean_inc(x_53); +lean_object* x_54; +x_54 = lean_ctor_get(x_1, 2); +lean_inc(x_54); lean_dec_ref(x_1); -x_1 = x_53; +x_1 = x_54; goto _start; } } case 7: { -uint8_t x_55; -x_55 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -if (x_55 == 0) +uint8_t x_56; +x_56 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +if (x_56 == 0) { -lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_56 = lean_ctor_get(x_1, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_1, 1); +lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_57 = lean_ctor_get(x_1, 0); 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); +x_58 = lean_ctor_get(x_1, 1); +lean_inc(x_58); +x_59 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_60 = lean_ctor_get(x_1, 2); +lean_inc(x_60); lean_dec_ref(x_1); -x_60 = l_Lean_IR_EmitC_emitDec___redArg(x_56, x_57, x_58, x_3); -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; +x_61 = l_Lean_IR_EmitC_emitDec___redArg(x_57, x_58, x_59, x_3); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec_ref(x_61); +x_1 = x_60; +x_3 = x_62; goto _start; } else { -lean_object* x_63; -x_63 = lean_ctor_get(x_1, 2); -lean_inc(x_63); +lean_object* x_64; +x_64 = lean_ctor_get(x_1, 2); +lean_inc(x_64); lean_dec_ref(x_1); -x_1 = x_63; +x_1 = x_64; goto _start; } } case 8: { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = lean_ctor_get(x_1, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_1, 1); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_1, 0); lean_inc(x_66); +x_67 = lean_ctor_get(x_1, 1); +lean_inc(x_67); lean_dec_ref(x_1); -x_67 = l_Lean_IR_EmitC_emitDel___redArg(x_65, x_3); -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; +x_68 = l_Lean_IR_EmitC_emitDel___redArg(x_66, x_3); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec_ref(x_68); +x_1 = x_67; +x_3 = x_69; goto _start; } 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_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = lean_ctor_get(x_1, 1); lean_inc(x_71); -x_72 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_72); +x_72 = lean_ctor_get(x_1, 2); +lean_inc(x_72); +x_73 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_73); 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; +x_74 = l_Lean_IR_EmitC_emitCase(x_71, x_72, x_73, x_2, x_3); +lean_dec(x_72); +return x_74; } case 10: { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_74 = lean_ctor_get(x_1, 0); -lean_inc(x_74); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_75 = lean_ctor_get(x_1, 0); +lean_inc(x_75); 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); -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 0) +x_76 = l_Lean_IR_EmitC_emitBlock___closed__0; +x_77 = lean_string_append(x_3, x_76); +x_78 = l_Lean_IR_EmitC_emitArg___redArg(x_75, x_77); +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 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___lam__0___closed__0; -x_82 = lean_string_append(x_79, x_81); -x_83 = l_Lean_IR_EmitC_emitLn___redArg___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; +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_80 = lean_ctor_get(x_78, 1); +x_81 = lean_ctor_get(x_78, 0); +lean_dec(x_81); +x_82 = l_Lean_IR_EmitC_emitFnDeclAux___lam__0___closed__0; +x_83 = lean_string_append(x_80, x_82); +x_84 = l_Lean_IR_EmitC_emitLn___redArg___closed__0; +x_85 = lean_box(0); +x_86 = lean_string_append(x_83, x_84); +lean_ctor_set(x_78, 1, x_86); +lean_ctor_set(x_78, 0, x_85); +return x_78; } 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___lam__0___closed__0; -x_88 = lean_string_append(x_86, x_87); -x_89 = l_Lean_IR_EmitC_emitLn___redArg___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_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_87 = lean_ctor_get(x_78, 1); +lean_inc(x_87); +lean_dec(x_78); +x_88 = l_Lean_IR_EmitC_emitFnDeclAux___lam__0___closed__0; +x_89 = lean_string_append(x_87, x_88); +x_90 = l_Lean_IR_EmitC_emitLn___redArg___closed__0; +x_91 = lean_box(0); +x_92 = lean_string_append(x_89, x_90); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } 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_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_1, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_95); lean_dec_ref(x_1); -x_95 = l_Lean_IR_EmitC_emitJmp(x_93, x_94, x_2, x_3); -lean_dec_ref(x_94); -return x_95; +x_96 = l_Lean_IR_EmitC_emitJmp(x_94, x_95, x_2, x_3); +lean_dec_ref(x_95); +return x_96; } 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; -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___redArg___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; +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_97 = l_Lean_IR_EmitC_emitBlock___closed__1; +x_98 = lean_string_append(x_3, x_97); +x_99 = l_Lean_IR_EmitC_emitLn___redArg___closed__0; +x_100 = lean_box(0); +x_101 = lean_string_append(x_98, x_99); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; } } } @@ -12310,11 +12641,11 @@ return x_26; if (x_34 == 0) { lean_object* x_35; lean_object* x_36; -lean_inc(x_33); -x_35 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_32, x_29, x_28, x_33, x_33, x_31); -lean_dec(x_33); -lean_dec_ref(x_29); -lean_dec_ref(x_32); +lean_inc(x_31); +x_35 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_28, x_33, x_32, x_31, x_31, x_29); +lean_dec(x_31); +lean_dec_ref(x_33); +lean_dec_ref(x_28); x_36 = lean_ctor_get(x_35, 1); lean_inc(x_36); lean_dec_ref(x_35); @@ -12325,11 +12656,11 @@ goto block_27; else { lean_object* x_37; lean_object* x_38; -lean_dec(x_33); -lean_dec_ref(x_32); -lean_dec_ref(x_29); +lean_dec_ref(x_33); +lean_dec(x_31); +lean_dec_ref(x_28); x_37 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_38 = lean_string_append(x_31, x_37); +x_38 = lean_string_append(x_29, x_37); x_21 = x_30; x_22 = x_38; goto block_27; @@ -12338,7 +12669,7 @@ goto block_27; block_52: { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_45 = lean_string_append(x_43, x_19); +x_45 = lean_string_append(x_41, x_19); lean_dec(x_19); x_46 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; x_47 = lean_string_append(x_45, x_46); @@ -12347,12 +12678,12 @@ x_49 = lean_array_get_size(x_44); x_50 = lean_nat_dec_lt(x_48, x_49); if (x_50 == 0) { -x_28 = x_40; -x_29 = x_41; -x_30 = x_42; -x_31 = x_47; -x_32 = x_44; -x_33 = x_49; +x_28 = x_44; +x_29 = x_47; +x_30 = x_40; +x_31 = x_49; +x_32 = x_42; +x_33 = x_43; x_34 = x_50; goto block_39; } @@ -12360,12 +12691,12 @@ else { uint8_t x_51; x_51 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_10); -x_28 = x_40; -x_29 = x_41; -x_30 = x_42; -x_31 = x_47; -x_32 = x_44; -x_33 = x_49; +x_28 = x_44; +x_29 = x_47; +x_30 = x_40; +x_31 = x_49; +x_32 = x_42; +x_33 = x_43; x_34 = x_51; goto block_39; } @@ -12405,10 +12736,10 @@ x_69 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; if (x_61 == 0) { lean_dec(x_60); -x_40 = x_59; -x_41 = x_57; -x_42 = x_53; -x_43 = x_58; +x_40 = x_53; +x_41 = x_58; +x_42 = x_59; +x_43 = x_57; x_44 = x_69; goto block_52; } @@ -12419,10 +12750,10 @@ x_70 = lean_nat_dec_le(x_60, x_60); if (x_70 == 0) { lean_dec(x_60); -x_40 = x_59; -x_41 = x_57; -x_42 = x_53; -x_43 = x_58; +x_40 = x_53; +x_41 = x_58; +x_42 = x_59; +x_43 = x_57; x_44 = x_69; goto block_52; } @@ -12433,10 +12764,10 @@ x_71 = 0; x_72 = lean_usize_of_nat(x_60); lean_dec(x_60); x_73 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_13, x_71, x_72, x_69); -x_40 = x_59; -x_41 = x_57; -x_42 = x_53; -x_43 = x_58; +x_40 = x_53; +x_41 = x_58; +x_42 = x_59; +x_43 = x_57; x_44 = x_73; goto block_52; } @@ -12538,11 +12869,11 @@ return x_99; if (x_107 == 0) { lean_object* x_108; lean_object* x_109; -lean_inc(x_106); -x_108 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_105, x_102, x_101, x_106, x_106, x_104); -lean_dec(x_106); -lean_dec_ref(x_102); -lean_dec_ref(x_105); +lean_inc(x_104); +x_108 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_101, x_106, x_105, x_104, x_104, x_102); +lean_dec(x_104); +lean_dec_ref(x_106); +lean_dec_ref(x_101); x_109 = lean_ctor_get(x_108, 1); lean_inc(x_109); lean_dec_ref(x_108); @@ -12553,11 +12884,11 @@ goto block_100; else { lean_object* x_110; lean_object* x_111; -lean_dec(x_106); -lean_dec_ref(x_105); -lean_dec_ref(x_102); +lean_dec_ref(x_106); +lean_dec(x_104); +lean_dec_ref(x_101); x_110 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_111 = lean_string_append(x_104, x_110); +x_111 = lean_string_append(x_102, x_110); x_94 = x_103; x_95 = x_111; goto block_100; @@ -12566,7 +12897,7 @@ goto block_100; block_125: { 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_118 = lean_string_append(x_116, x_92); +x_118 = lean_string_append(x_114, x_92); lean_dec(x_92); x_119 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; x_120 = lean_string_append(x_118, x_119); @@ -12575,12 +12906,12 @@ x_122 = lean_array_get_size(x_117); x_123 = lean_nat_dec_lt(x_121, x_122); if (x_123 == 0) { -x_101 = x_113; -x_102 = x_114; -x_103 = x_115; -x_104 = x_120; -x_105 = x_117; -x_106 = x_122; +x_101 = x_117; +x_102 = x_120; +x_103 = x_113; +x_104 = x_122; +x_105 = x_115; +x_106 = x_116; x_107 = x_123; goto block_112; } @@ -12588,12 +12919,12 @@ else { uint8_t x_124; x_124 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_10); -x_101 = x_113; -x_102 = x_114; -x_103 = x_115; -x_104 = x_120; -x_105 = x_117; -x_106 = x_122; +x_101 = x_117; +x_102 = x_120; +x_103 = x_113; +x_104 = x_122; +x_105 = x_115; +x_106 = x_116; x_107 = x_124; goto block_112; } @@ -12633,10 +12964,10 @@ x_142 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; if (x_134 == 0) { lean_dec(x_133); -x_113 = x_132; -x_114 = x_130; -x_115 = x_126; -x_116 = x_131; +x_113 = x_126; +x_114 = x_131; +x_115 = x_132; +x_116 = x_130; x_117 = x_142; goto block_125; } @@ -12647,10 +12978,10 @@ x_143 = lean_nat_dec_le(x_133, x_133); if (x_143 == 0) { lean_dec(x_133); -x_113 = x_132; -x_114 = x_130; -x_115 = x_126; -x_116 = x_131; +x_113 = x_126; +x_114 = x_131; +x_115 = x_132; +x_116 = x_130; x_117 = x_142; goto block_125; } @@ -12661,10 +12992,10 @@ x_144 = 0; x_145 = lean_usize_of_nat(x_133); lean_dec(x_133); x_146 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_13, x_144, x_145, x_142); -x_113 = x_132; -x_114 = x_130; -x_115 = x_126; -x_116 = x_131; +x_113 = x_126; +x_114 = x_131; +x_115 = x_132; +x_116 = x_130; x_117 = x_146; goto block_125; } @@ -12834,11 +13165,11 @@ return x_185; if (x_193 == 0) { lean_object* x_194; lean_object* x_195; -lean_inc(x_192); -x_194 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_191, x_188, x_187, x_192, x_192, x_190); -lean_dec(x_192); -lean_dec_ref(x_188); -lean_dec_ref(x_191); +lean_inc(x_190); +x_194 = l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitC_emitDeclAux_spec__1___redArg(x_187, x_192, x_191, x_190, x_190, x_188); +lean_dec(x_190); +lean_dec_ref(x_192); +lean_dec_ref(x_187); x_195 = lean_ctor_get(x_194, 1); lean_inc(x_195); lean_dec_ref(x_194); @@ -12849,11 +13180,11 @@ goto block_186; else { lean_object* x_196; lean_object* x_197; -lean_dec(x_192); -lean_dec_ref(x_191); -lean_dec_ref(x_188); +lean_dec_ref(x_192); +lean_dec(x_190); +lean_dec_ref(x_187); x_196 = l_Lean_IR_EmitC_emitDeclAux___closed__0; -x_197 = lean_string_append(x_190, x_196); +x_197 = lean_string_append(x_188, x_196); x_180 = x_189; x_181 = x_197; goto block_186; @@ -12862,7 +13193,7 @@ goto block_186; block_211: { lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_204 = lean_string_append(x_202, x_178); +x_204 = lean_string_append(x_200, x_178); lean_dec(x_178); x_205 = l_Lean_IR_EmitC_emitFnDeclAux___closed__4; x_206 = lean_string_append(x_204, x_205); @@ -12871,12 +13202,12 @@ x_208 = lean_array_get_size(x_203); x_209 = lean_nat_dec_lt(x_207, x_208); if (x_209 == 0) { -x_187 = x_199; -x_188 = x_200; -x_189 = x_201; -x_190 = x_206; -x_191 = x_203; -x_192 = x_208; +x_187 = x_203; +x_188 = x_206; +x_189 = x_199; +x_190 = x_208; +x_191 = x_201; +x_192 = x_202; x_193 = x_209; goto block_198; } @@ -12884,12 +13215,12 @@ else { uint8_t x_210; x_210 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_165); -x_187 = x_199; -x_188 = x_200; -x_189 = x_201; -x_190 = x_206; -x_191 = x_203; -x_192 = x_208; +x_187 = x_203; +x_188 = x_206; +x_189 = x_199; +x_190 = x_208; +x_191 = x_201; +x_192 = x_202; x_193 = x_210; goto block_198; } @@ -12929,10 +13260,10 @@ x_228 = l_Lean_IR_EmitC_emitFnDeclAux___closed__5; if (x_220 == 0) { lean_dec(x_219); -x_199 = x_218; -x_200 = x_216; -x_201 = x_212; -x_202 = x_217; +x_199 = x_212; +x_200 = x_217; +x_201 = x_218; +x_202 = x_216; x_203 = x_228; goto block_211; } @@ -12943,10 +13274,10 @@ x_229 = lean_nat_dec_le(x_219, x_219); if (x_229 == 0) { lean_dec(x_219); -x_199 = x_218; -x_200 = x_216; -x_201 = x_212; -x_202 = x_217; +x_199 = x_212; +x_200 = x_217; +x_201 = x_218; +x_202 = x_216; x_203 = x_228; goto block_211; } @@ -12957,10 +13288,10 @@ x_230 = 0; x_231 = lean_usize_of_nat(x_219); lean_dec(x_219); x_232 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_EmitC_emitFnDeclAux_spec__2(x_168, x_230, x_231, x_228); -x_199 = x_218; -x_200 = x_216; -x_201 = x_212; -x_202 = x_217; +x_199 = x_212; +x_200 = x_217; +x_201 = x_218; +x_202 = x_216; x_203 = x_232; goto block_211; } @@ -14853,6 +15184,16 @@ l_Lean_IR_EmitC_toCType___closed__11 = _init_l_Lean_IR_EmitC_toCType___closed__1 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_toCType___closed__13 = _init_l_Lean_IR_EmitC_toCType___closed__13(); +lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__13); +l_Lean_IR_EmitC_toCType___closed__14 = _init_l_Lean_IR_EmitC_toCType___closed__14(); +lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__14); +l_Lean_IR_EmitC_toCType___closed__15 = _init_l_Lean_IR_EmitC_toCType___closed__15(); +lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__15); +l_Lean_IR_EmitC_toCType___closed__16 = _init_l_Lean_IR_EmitC_toCType___closed__16(); +lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__16); +l_Lean_IR_EmitC_toCType___closed__17 = _init_l_Lean_IR_EmitC_toCType___closed__17(); +lean_mark_persistent(l_Lean_IR_EmitC_toCType___closed__17); 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(); @@ -15103,6 +15444,8 @@ l_Lean_IR_EmitC_emitOffset___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitOffs 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_emitUSet___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitUSet___redArg___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUSet___redArg___closed__1); 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(); @@ -15117,6 +15460,14 @@ l_Lean_IR_EmitC_emitSSet___redArg___closed__5 = _init_l_Lean_IR_EmitC_emitSSet__ 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_Lean_IR_EmitC_emitSSet___redArg___closed__7 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__7); +l_Lean_IR_EmitC_emitSSet___redArg___closed__8 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__8); +l_Lean_IR_EmitC_emitSSet___redArg___closed__9 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__9); +l_Lean_IR_EmitC_emitSSet___redArg___closed__10 = _init_l_Lean_IR_EmitC_emitSSet___redArg___closed__10(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSSet___redArg___closed__10); 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_Lean_IR_EmitC_emitJmp___closed__0 = _init_l_Lean_IR_EmitC_emitJmp___closed__0(); @@ -15149,6 +15500,8 @@ l_Lean_IR_EmitC_emitProj___redArg___closed__0 = _init_l_Lean_IR_EmitC_emitProj__ 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_emitUProj___redArg___closed__1 = _init_l_Lean_IR_EmitC_emitUProj___redArg___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitC_emitUProj___redArg___closed__1); 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(); @@ -15161,6 +15514,14 @@ l_Lean_IR_EmitC_emitSProj___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitSProj 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_emitSProj___redArg___closed__6 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__6); +l_Lean_IR_EmitC_emitSProj___redArg___closed__7 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__7); +l_Lean_IR_EmitC_emitSProj___redArg___closed__8 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__8); +l_Lean_IR_EmitC_emitSProj___redArg___closed__9 = _init_l_Lean_IR_EmitC_emitSProj___redArg___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitC_emitSProj___redArg___closed__9); 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_emitExternCall___closed__0 = _init_l_Lean_IR_EmitC_emitExternCall___closed__0(); @@ -15195,6 +15556,12 @@ l_Lean_IR_EmitC_emitBoxFn___redArg___closed__4 = _init_l_Lean_IR_EmitC_emitBoxFn 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___redArg___closed__6 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__6); +l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__7); +l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8 = _init_l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitC_emitBoxFn___redArg___closed__8); 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_String_foldlAux___at___00Lean_IR_EmitC_quoteString_spec__0___closed__0 = _init_l_String_foldlAux___at___00Lean_IR_EmitC_quoteString_spec__0___closed__0(); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c index 923c8938cc58..e087533696a5 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c @@ -107,6 +107,7 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_ static lean_object* l_Lean_IR_EmitLLVM_emitTailCall___closed__1; LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitLLVM_emitFns_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultShowError___redArg___closed__0; +static lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitMainFn___lam__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__1; size_t lean_llvm_build_sext_or_trunc(size_t, size_t, size_t, size_t, lean_object*); @@ -114,6 +115,7 @@ size_t lean_llvm_build_switch(size_t, size_t, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_EmitLLVM_addJpToState_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_instBEqVarId_beq(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__12; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_buildLeanBoolTrue_x3f___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanUnsignedToNatFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -142,6 +144,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___00Lean_IR_EmitLLVM_emitFns_spec__0(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_dec_elim___redArg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitTailCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitBox___closed__7; static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___00Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitLLVM_emitLhsSlot___00spec__0_spec__0(lean_object*, lean_object*, lean_object*); size_t l_LLVM_i8Type(size_t); @@ -188,6 +191,7 @@ uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorSetTag___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__11; lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitArgSlot___00__boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_IR_EmitLLVM_toLLVMType_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -281,6 +285,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFns(size_t, size_t, size_t, lean LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCStrToNatFn(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_buildIfThenElse___00__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_constInt64(size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__8; static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__5; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getFunIdTy___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_resultType(lean_object*); @@ -318,7 +323,7 @@ size_t lean_llvm_build_global_string(size_t, size_t, lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitExternCall(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_EmitLLVM_addVarToState_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet(size_t, size_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); size_t lean_llvm_array_type(size_t, size_t, uint64_t); static lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultIsError___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_throwInvalidExportName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -377,6 +382,7 @@ size_t lean_llvm_get_undef(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultShowError___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanAllocClosureFn___redArg(size_t, size_t, size_t, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_ensureHasDefault_x27___closed__0; +static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__7; static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__2; lean_object* lean_nat_div(lean_object*, lean_object*); size_t lean_llvm_create_context(); @@ -405,6 +411,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitReset___lam__1(size_t, lean_obje LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIsScalar(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constInt8(size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constInt64___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0; uint8_t l_Lean_isExternC(lean_object*, lean_object*); lean_object* lean_llvm_get_first_instruction(size_t, size_t); LEAN_EXPORT lean_object* l_String_foldlAux___at___00Lean_IR_EmitLLVM_quoteString_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -440,7 +447,6 @@ LEAN_EXPORT size_t l_Lean_IR_LLVM_unsignedType(size_t); size_t lean_llvm_add_global(size_t, size_t, lean_object*, size_t); static lean_object* l_Lean_IR_EmitLLVM_declareVar___closed__0; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_IR_EmitLLVM_addVarToState_spec__0_spec__0___redArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize(size_t, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); lean_object* lean_llvm_verify_module(size_t, size_t); uint8_t l_Lean_IR_IRType_isObj(lean_object*); @@ -464,6 +470,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getOrAddFunIdValue(size_t, size_t, l static lean_object* l_Array_filterMapM___at___00Lean_IR_emitLLVM_spec__1___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGet(size_t, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIsExclusive___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType(size_t, size_t, size_t, size_t, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitFnDeclAux_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_EmitLLVM_callLeanCtorSetTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_IR_EmitLLVM_emitJp_spec__0___redArg(lean_object*, lean_object*); @@ -525,6 +532,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_ShouldForwardControlFlow_ctorIdx___b LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitSimpleExternalCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_llvm_opaque_pointer_type_in_context(size_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___boxed(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___00Lean_IR_EmitLLVM_addJpToState_spec__0_spec__0___redArg(lean_object*, lean_object*); extern uint64_t l_LLVM_IntPredicate_NE; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitSSet(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -548,6 +556,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_dec_elim___redArg(lean_ LEAN_EXPORT lean_object* l_List_foldl___at___00Lean_IR_EmitLLVM_emitFnDecls_spec__1(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_EmitLLVM_addVarToState_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_instToStringRefcountKind___lam__0___closed__0; +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitOffset(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_callLeanIsExclusive___redArg___closed__0; @@ -600,6 +609,7 @@ lean_object* l_Lean_NameSet_insert(lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIsScalar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_builderAppendBasicBlock___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_callUnboxForType___closed__8; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_buildIfThen___00__closed__1; @@ -643,7 +653,7 @@ static lean_object* l_Lean_IR_EmitLLVM_buildWhile___00__closed__1; static lean_object* l_Lean_IR_EmitLLVM_emitUnreachable___redArg___closed__0; LEAN_EXPORT uint8_t l_Lean_IR_EmitLLVM_IRType_isIntegerType(lean_object*); lean_object* lean_llvm_set_initializer(size_t, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet___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_emitUSet___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_IR_EmitLLVM_emitApp_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_EmitLLVM_buildPrologueAlloca___redArg___closed__0; static lean_object* l_Lean_IR_EmitLLVM_emitInitFn___lam__0___closed__0; @@ -657,15 +667,15 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constIntUnsigned___redArg___boxed(le static lean_object* l_Lean_IR_EmitLLVM_emitJmp___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callIODeclInitFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitMainFn___closed__4; +static lean_object* l_Lean_IR_EmitLLVM_callUnboxForType___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIsScalar___redArg(size_t, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitExternCall___closed__7; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitLLVM_emitCtorSetArgs_spec__0___redArg(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_callUnboxForType___closed__4; -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUProj___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_emitUProj___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_Compiler_IR_EmitLLVM_0__Lean_IR_getModuleFunctions_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFnDecls___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_LLVM_unsignedType___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(size_t, size_t, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_LLVM_getOrAddFunction___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Control_0__Nat_forM_loop___at___00Lean_IR_EmitLLVM_emitReset_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_EmitLLVM_instInhabitedM___closed__0; @@ -703,14 +713,12 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_toLLVMType(size_t, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanUnboxUint32(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitUSet___closed__0; static lean_object* l_Lean_IR_EmitLLVM_callLeanInitTaskManager___redArg___closed__0; -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_llvm_build_add(size_t, size_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitJp___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultMKOk___redArg(size_t, size_t, size_t, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitCase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern uint64_t l_LLVM_DLLStorageClass_export; -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitDecl___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00Lean_IR_EmitLLVM_emitFnDecls_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_String_foldlAux___at___00Lean_IR_EmitLLVM_quoteString_spec__0___closed__3; @@ -755,6 +763,7 @@ static lean_object* l_Lean_IR_EmitLLVM_emitArgSlot___00__closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_buildIfThen__(size_t, size_t, lean_object*, size_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_getDecls(lean_object*); uint8_t l_Lean_IR_isTailCallTo(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFnArgs(size_t, size_t, uint8_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIsExclusive(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitSimpleExternalCall_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -784,6 +793,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUnreachable___redArg(size_t, siz static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__4; static lean_object* l_Lean_IR_EmitLLVM_buildIfThen___00__closed__0; lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(size_t, size_t, size_t, size_t, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_getDecl(size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callIODeclInitFn___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___redArg(lean_object*); @@ -832,6 +842,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorSetTag(size_t, size_t, s LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_LLVM_Value_isNull___redArg(size_t); size_t lean_llvm_build_icmp(size_t, size_t, uint64_t, size_t, size_t, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitBox___closed__5; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitLhsSlot___00__redArg(lean_object*, lean_object*); size_t lean_llvm_build_alloca(size_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanInitializeRuntimeModule(size_t, size_t, lean_object*, lean_object*); @@ -841,6 +852,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitExternCall___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constInt64___redArg___boxed(lean_object*, lean_object*, lean_object*); size_t lean_llvm_get_first_global(size_t, size_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_EmitLLVM_addVarToState_spec__0_spec__1_spec__1___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__10; static lean_object* l_Lean_IR_EmitLLVM_callUnboxForType___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitInc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_uint64_dec_eq(uint64_t, uint64_t); @@ -873,6 +885,7 @@ lean_object* l_mkPanicMessageWithDecl(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_EmitLLVM_emitJp_spec__0_spec__0(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_buildPrologueAlloca___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitBox___closed__6; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanBox___redArg(size_t, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_ShouldForwardControlFlow_ctorElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_ShouldForwardControlFlow_no_elim(lean_object*, uint8_t, lean_object*, lean_object*); @@ -882,6 +895,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanIOResultShowError___redArg__ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__3; size_t lean_array_size(lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitUSet___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitLLVM_getFunIdTy_spec__0(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); size_t l_LLVM_i1Type(size_t); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanInitTaskManager___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -905,6 +919,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitFnBody(size_t, size_t, lean_obje LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_toHexDigit(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_toCInitName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitMainFn___closed__7; +static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__9; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_noConfusion___redArg___lam__0(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_emitLLVM_spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_llvm_void_type_in_context(size_t); @@ -949,7 +964,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanMkStringUncheckedFn(size_t, LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanMainFn___redArg___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_EmitLLVM_emitLhsSlot___00spec__0___redArg(lean_object*, lean_object*); uint8_t l_Lean_isIOUnitInitFn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUProj(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUProj(size_t, size_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callModInitFn___redArg(size_t, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_IR_leanMainFn___closed__0; @@ -1013,6 +1028,7 @@ lean_object* l_Array_zip___redArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanObjTag(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_RefcountKind_inc_elim___redArg___boxed(lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitBox___closed__4; +static lean_object* l_Lean_IR_EmitLLVM_callUnboxForType___closed__7; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanDecRef___redArg(size_t, size_t, size_t, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitSProj___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitLhsSlotStore(size_t, size_t, lean_object*, size_t, lean_object*, lean_object*); @@ -1026,8 +1042,8 @@ uint8_t l_Lean_IR_IRType_isScalar(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_constInt8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitLLVM_callLeanClosureSetFn___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanRefcountFn___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_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitLLVM_emitApp_spec__0(size_t, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitLLVM_emitSSet___closed__10; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanAllocClosureFn___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_IR_EmitLLVM_emitFnArgs_spec__0(lean_object*, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_declareVars_spec__0(size_t, size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -6328,7 +6344,7 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_toLLVMType___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_EmitLLVM_toLLVMType___closed__2; x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(341u); +x_3 = lean_unsigned_to_nat(349u); x_4 = l_Lean_IR_EmitLLVM_toLLVMType___closed__1; x_5 = l_Lean_IR_EmitLLVM_toLLVMType___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -6341,7 +6357,7 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_toLLVMType___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_EmitLLVM_toLLVMType___closed__2; x_2 = lean_unsigned_to_nat(25u); -x_3 = lean_unsigned_to_nat(342u); +x_3 = lean_unsigned_to_nat(350u); x_4 = l_Lean_IR_EmitLLVM_toLLVMType___closed__1; x_5 = l_Lean_IR_EmitLLVM_toLLVMType___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -6351,130 +6367,191 @@ return x_6; LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_toLLVMType(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { +lean_object* x_6; lean_object* x_13; lean_object* x_20; lean_object* x_27; lean_object* x_34; switch (lean_obj_tag(x_2)) { case 0: { -size_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_dec_ref(x_4); lean_dec(x_3); -x_6 = lean_llvm_double_type_in_context(x_1); -x_7 = lean_box_usize(x_6); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_8); -return x_9; +x_40 = lean_llvm_double_type_in_context(x_1); +x_41 = lean_box_usize(x_40); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_43, 0, x_42); +return x_43; } case 1: { -uint64_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec_ref(x_4); lean_dec(x_3); -x_10 = 8; -x_11 = lean_llvm_int_type_in_context(x_1, x_10); -x_12 = lean_box_usize(x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; +x_6 = lean_box(0); +goto block_12; } case 2: { -uint64_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec_ref(x_4); lean_dec(x_3); -x_15 = 16; -x_16 = lean_llvm_int_type_in_context(x_1, x_15); -x_17 = lean_box_usize(x_16); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_17); -x_19 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_19, 0, x_18); -return x_19; +x_13 = lean_box(0); +goto block_19; } case 3: { -uint64_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec_ref(x_4); lean_dec(x_3); -x_20 = 32; -x_21 = lean_llvm_int_type_in_context(x_1, x_20); -x_22 = lean_box_usize(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; +x_20 = lean_box(0); +goto block_26; } case 4: { -uint64_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec_ref(x_4); lean_dec(x_3); -x_25 = 64; -x_26 = lean_llvm_int_type_in_context(x_1, x_25); -x_27 = lean_box_usize(x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_29, 0, x_28); -return x_29; +x_27 = lean_box(0); +goto block_33; } case 5: { -size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec_ref(x_4); lean_dec(x_3); -x_30 = l_LLVM_i64Type(x_1); -x_31 = lean_box_usize(x_30); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_32); -return x_33; +x_34 = lean_box(0); +goto block_39; } case 9: { -size_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec_ref(x_4); lean_dec(x_3); -x_34 = lean_llvm_float_type_in_context(x_1); -x_35 = lean_box_usize(x_34); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +x_44 = lean_llvm_float_type_in_context(x_1); +x_45 = lean_box_usize(x_44); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_47, 0, x_46); +return x_47; } case 10: { -lean_object* x_38; lean_object* x_39; -x_38 = l_Lean_IR_EmitLLVM_toLLVMType___closed__3; -x_39 = l_panic___at___00Lean_IR_EmitLLVM_toLLVMType_spec__0(x_1, x_38, x_3, x_4); -return x_39; +lean_object* x_48; lean_object* x_49; +x_48 = l_Lean_IR_EmitLLVM_toLLVMType___closed__3; +x_49 = l_panic___at___00Lean_IR_EmitLLVM_toLLVMType_spec__0(x_1, x_48, x_3, x_4); +return x_49; } case 11: { -lean_object* x_40; lean_object* x_41; -x_40 = l_Lean_IR_EmitLLVM_toLLVMType___closed__4; -x_41 = l_panic___at___00Lean_IR_EmitLLVM_toLLVMType_spec__0(x_1, x_40, x_3, x_4); -return x_41; +lean_object* x_50; lean_object* x_51; +x_50 = l_Lean_IR_EmitLLVM_toLLVMType___closed__4; +x_51 = l_panic___at___00Lean_IR_EmitLLVM_toLLVMType_spec__0(x_1, x_50, x_3, x_4); +return x_51; +} +case 14: +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_6 = lean_box(0); +goto block_12; +} +case 15: +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_13 = lean_box(0); +goto block_19; +} +case 16: +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_20 = lean_box(0); +goto block_26; +} +case 17: +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_27 = lean_box(0); +goto block_33; +} +case 18: +{ +lean_dec_ref(x_4); +lean_dec(x_3); +x_34 = lean_box(0); +goto block_39; } default: { -size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_dec_ref(x_4); lean_dec(x_3); -x_42 = l_LLVM_i8Type(x_1); -x_43 = lean_llvm_pointer_type(x_1, x_42); -x_44 = lean_box_usize(x_43); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_46, 0, x_45); -return x_46; +x_52 = l_LLVM_i8Type(x_1); +x_53 = lean_llvm_pointer_type(x_1, x_52); +x_54 = lean_box_usize(x_53); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +return x_56; +} } +block_12: +{ +uint64_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = 8; +x_8 = lean_llvm_int_type_in_context(x_1, x_7); +x_9 = lean_box_usize(x_8); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +block_19: +{ +uint64_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = 16; +x_15 = lean_llvm_int_type_in_context(x_1, x_14); +x_16 = lean_box_usize(x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_17); +return x_18; +} +block_26: +{ +uint64_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = 32; +x_22 = lean_llvm_int_type_in_context(x_1, x_21); +x_23 = lean_box_usize(x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +return x_25; +} +block_33: +{ +uint64_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = 64; +x_29 = lean_llvm_int_type_in_context(x_1, x_28); +x_30 = lean_box_usize(x_29); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; +} +block_39: +{ +size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = l_LLVM_i64Type(x_1); +x_36 = lean_box_usize(x_35); +x_37 = lean_alloc_ctor(1, 1, 0); +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; } } } @@ -13930,7 +14007,7 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_emitExternCall___closed__8() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_IR_EmitLLVM_emitExternCall___closed__7; x_2 = lean_unsigned_to_nat(36u); -x_3 = lean_unsigned_to_nat(682u); +x_3 = lean_unsigned_to_nat(690u); x_4 = l_Lean_IR_EmitLLVM_emitExternCall___closed__6; x_5 = l_Lean_IR_EmitLLVM_toLLVMType___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -22549,7 +22626,7 @@ lean_dec(x_4); return x_11; } } -static lean_object* _init_l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0() { +static lean_object* _init_l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0() { _start: { lean_object* x_1; @@ -22557,291 +22634,320 @@ x_1 = lean_mk_string_unchecked("lean_ctor_get_usize", 19, 19); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(size_t x_1, size_t x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1() { _start: { -size_t x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; uint8_t x_24; -x_8 = l_LLVM_i64Type(x_1); -x_9 = l_LLVM_voidPtrType(x_1); -x_10 = l_LLVM_i32Type(x_1); -x_11 = l_Lean_IR_EmitLLVM_callLeanRefcountFn___redArg___closed__0; -x_12 = lean_box_usize(x_9); -x_13 = lean_array_push(x_11, x_12); -x_14 = lean_box_usize(x_10); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_isize", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(size_t x_1, size_t 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_9; +if (x_5 == 0) +{ +lean_object* x_58; +x_58 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0; +x_9 = x_58; +goto block_57; +} +else +{ +lean_object* x_59; +x_59 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1; +x_9 = x_59; +goto block_57; +} +block_57: +{ +size_t x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; uint8_t x_25; +x_10 = l_LLVM_i64Type(x_1); +x_11 = l_LLVM_voidPtrType(x_1); +x_12 = l_LLVM_i32Type(x_1); +x_13 = l_Lean_IR_EmitLLVM_callLeanRefcountFn___redArg___closed__0; +x_14 = lean_box_usize(x_11); x_15 = lean_array_push(x_13, x_14); -x_16 = 0; -x_17 = lean_llvm_function_type(x_1, x_8, x_15, x_16); -x_18 = l_Lean_IR_EmitLLVM_getLLVMModule___redArg(x_6); -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(x_20); -lean_dec(x_19); -x_21 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0; -x_22 = lean_unbox_usize(x_20); -lean_dec(x_20); -x_23 = l_Lean_IR_EmitLLVM_getOrCreateFunctionPrototype___redArg(x_1, x_22, x_8, x_21, x_15); -lean_dec_ref(x_15); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_16 = lean_box_usize(x_12); +x_17 = lean_array_push(x_15, x_16); +x_18 = 0; +x_19 = lean_llvm_function_type(x_1, x_10, x_17, x_18); +x_20 = l_Lean_IR_EmitLLVM_getLLVMModule___redArg(x_7); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec_ref(x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_unbox_usize(x_22); +lean_dec(x_22); +x_24 = l_Lean_IR_EmitLLVM_getOrCreateFunctionPrototype___redArg(x_1, x_23, x_10, x_9, x_17); +lean_dec_ref(x_17); +lean_dec_ref(x_9); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_23, 0); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_24, 0); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; lean_object* x_34; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_box_usize(x_3); -x_29 = lean_array_push(x_11, x_28); -x_30 = lean_box_usize(x_4); -x_31 = lean_array_push(x_29, x_30); -x_32 = lean_unbox_usize(x_27); -lean_dec(x_27); -x_33 = lean_llvm_build_call2(x_1, x_2, x_17, x_32, x_31, x_5); -lean_dec_ref(x_31); -x_34 = lean_box_usize(x_33); -lean_ctor_set(x_25, 0, x_34); -return x_23; +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_object* x_35; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_box_usize(x_3); +x_30 = lean_array_push(x_13, x_29); +x_31 = lean_box_usize(x_4); +x_32 = lean_array_push(x_30, x_31); +x_33 = lean_unbox_usize(x_28); +lean_dec(x_28); +x_34 = lean_llvm_build_call2(x_1, x_2, x_19, x_33, x_32, x_6); +lean_dec_ref(x_32); +x_35 = lean_box_usize(x_34); +lean_ctor_set(x_26, 0, x_35); +return x_24; } else { -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; -x_35 = lean_ctor_get(x_25, 0); -lean_inc(x_35); -lean_dec(x_25); -x_36 = lean_box_usize(x_3); -x_37 = lean_array_push(x_11, x_36); -x_38 = lean_box_usize(x_4); -x_39 = lean_array_push(x_37, x_38); -x_40 = lean_unbox_usize(x_35); -lean_dec(x_35); -x_41 = lean_llvm_build_call2(x_1, x_2, x_17, x_40, x_39, x_5); -lean_dec_ref(x_39); -x_42 = lean_box_usize(x_41); -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_23, 0, x_43); -return x_23; +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; +x_36 = lean_ctor_get(x_26, 0); +lean_inc(x_36); +lean_dec(x_26); +x_37 = lean_box_usize(x_3); +x_38 = lean_array_push(x_13, x_37); +x_39 = lean_box_usize(x_4); +x_40 = lean_array_push(x_38, x_39); +x_41 = lean_unbox_usize(x_36); +lean_dec(x_36); +x_42 = lean_llvm_build_call2(x_1, x_2, x_19, x_41, x_40, x_6); +lean_dec_ref(x_40); +x_43 = lean_box_usize(x_42); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_24, 0, x_44); +return x_24; } } 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; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_44 = lean_ctor_get(x_23, 0); -lean_inc(x_44); -lean_dec(x_23); -x_45 = lean_ctor_get(x_44, 0); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_45 = lean_ctor_get(x_24, 0); lean_inc(x_45); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - x_46 = x_44; +lean_dec(x_24); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + x_47 = x_45; } else { - lean_dec_ref(x_44); - x_46 = lean_box(0); + lean_dec_ref(x_45); + x_47 = lean_box(0); } -x_47 = lean_box_usize(x_3); -x_48 = lean_array_push(x_11, x_47); -x_49 = lean_box_usize(x_4); -x_50 = lean_array_push(x_48, x_49); -x_51 = lean_unbox_usize(x_45); -lean_dec(x_45); -x_52 = lean_llvm_build_call2(x_1, x_2, x_17, x_51, x_50, x_5); -lean_dec_ref(x_50); -x_53 = lean_box_usize(x_52); -if (lean_is_scalar(x_46)) { - x_54 = lean_alloc_ctor(1, 1, 0); +x_48 = lean_box_usize(x_3); +x_49 = lean_array_push(x_13, x_48); +x_50 = lean_box_usize(x_4); +x_51 = lean_array_push(x_49, x_50); +x_52 = lean_unbox_usize(x_46); +lean_dec(x_46); +x_53 = lean_llvm_build_call2(x_1, x_2, x_19, x_52, x_51, x_6); +lean_dec_ref(x_51); +x_54 = lean_box_usize(x_53); +if (lean_is_scalar(x_47)) { + x_55 = lean_alloc_ctor(1, 1, 0); } else { - x_54 = x_46; + x_55 = x_47; } -lean_ctor_set(x_54, 0, x_53); -x_55 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_55, 0, x_54); -return x_55; +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +return x_56; +} } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize(size_t x_1, size_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_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType(size_t x_1, size_t x_2, size_t x_3, size_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_9; -x_9 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(x_1, x_2, x_3, x_4, x_5, x_7); -return x_9; +lean_object* x_10; +x_10 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(x_1, x_2, x_3, x_4, x_5, x_6, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___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_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___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_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_8 = lean_unbox_usize(x_1); +size_t x_9; size_t x_10; size_t x_11; size_t x_12; uint8_t x_13; lean_object* x_14; +x_9 = lean_unbox_usize(x_1); lean_dec(x_1); -x_9 = lean_unbox_usize(x_2); +x_10 = lean_unbox_usize(x_2); lean_dec(x_2); -x_10 = lean_unbox_usize(x_3); +x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -x_11 = lean_unbox_usize(x_4); +x_12 = lean_unbox_usize(x_4); lean_dec(x_4); -x_12 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(x_8, x_9, x_10, x_11, x_5, x_6); -lean_dec_ref(x_6); -return x_12; +x_13 = lean_unbox(x_5); +x_14 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(x_9, x_10, x_11, x_12, x_13, x_6, x_7); +lean_dec_ref(x_7); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___boxed(lean_object* x_1, lean_object* x_2, lean_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_callLeanCtorGetPtrSizedType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_9; size_t x_10; size_t x_11; size_t x_12; lean_object* x_13; -x_9 = lean_unbox_usize(x_1); +size_t x_10; size_t x_11; size_t x_12; size_t x_13; uint8_t x_14; lean_object* x_15; +x_10 = lean_unbox_usize(x_1); lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); +x_11 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = lean_unbox_usize(x_3); +x_12 = lean_unbox_usize(x_3); lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); +x_13 = lean_unbox_usize(x_4); lean_dec(x_4); -x_13 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize(x_9, x_10, x_11, x_12, x_5, x_6, x_7); -lean_dec_ref(x_7); -lean_dec(x_6); -return x_13; +x_14 = lean_unbox(x_5); +x_15 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType(x_10, x_11, x_12, x_13, x_14, x_6, x_7, x_8); +lean_dec_ref(x_8); +lean_dec(x_7); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUProj(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_emitUProj(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, lean_object* x_8) { _start: { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; -x_10 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_5, x_9, x_6); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; +x_11 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_6, x_10, x_7); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -if (lean_obj_tag(x_12) == 0) +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_13; +uint8_t x_14; lean_dec(x_3); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -return x_10; +return x_11; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_10, 0, x_15); -return x_10; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; -lean_free_object(x_10); -x_16 = lean_ctor_get(x_12, 0); -lean_inc(x_16); -lean_dec_ref(x_12); -x_17 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; lean_object* x_27; +lean_free_object(x_11); +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec_ref(x_13); +x_18 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_unbox_usize(x_16); -lean_dec(x_16); -x_21 = lean_unbox_usize(x_19); +lean_dec_ref(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); lean_dec(x_19); -x_22 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(x_1, x_2, x_20, x_21, x_9, x_7); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec_ref(x_22); +x_21 = lean_unbox_usize(x_17); +lean_dec(x_17); +x_22 = lean_unbox_usize(x_20); +lean_dec(x_20); +x_23 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(x_1, x_2, x_21, x_22, x_5, x_10, x_8); x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_unbox_usize(x_24); +lean_dec_ref(x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); lean_dec(x_24); -x_26 = l_Lean_IR_EmitLLVM_emitLhsSlotStore___redArg(x_1, x_2, x_3, x_25, x_6); -return x_26; +x_26 = lean_unbox_usize(x_25); +lean_dec(x_25); +x_27 = l_Lean_IR_EmitLLVM_emitLhsSlotStore___redArg(x_1, x_2, x_3, x_26, x_7); +return x_27; } } else { -lean_object* x_27; -x_27 = lean_ctor_get(x_10, 0); -lean_inc(x_27); -lean_dec(x_10); -if (lean_obj_tag(x_27) == 0) +lean_object* x_28; +x_28 = lean_ctor_get(x_11, 0); +lean_inc(x_28); +lean_dec(x_11); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_3); -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; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + x_30 = x_28; } else { - lean_dec_ref(x_27); - x_29 = lean_box(0); + lean_dec_ref(x_28); + x_30 = lean_box(0); } -if (lean_is_scalar(x_29)) { - x_30 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(0, 1, 0); } else { - x_30 = x_29; + x_31 = x_30; } -lean_ctor_set(x_30, 0, x_28); -x_31 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_31, 0, x_30); -return x_31; +lean_ctor_set(x_31, 0, x_29); +x_32 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; lean_object* x_42; -x_32 = lean_ctor_get(x_27, 0); -lean_inc(x_32); -lean_dec_ref(x_27); -x_33 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -lean_dec_ref(x_33); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +lean_dec_ref(x_28); +x_34 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_34); -x_36 = lean_unbox_usize(x_32); -lean_dec(x_32); -x_37 = lean_unbox_usize(x_35); +lean_dec_ref(x_34); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); lean_dec(x_35); -x_38 = l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg(x_1, x_2, x_36, x_37, x_9, x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -lean_dec_ref(x_38); +x_37 = lean_unbox_usize(x_33); +lean_dec(x_33); +x_38 = lean_unbox_usize(x_36); +lean_dec(x_36); +x_39 = l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg(x_1, x_2, x_37, x_38, x_5, x_10, x_8); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_39); -x_41 = lean_unbox_usize(x_40); +lean_dec_ref(x_39); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); lean_dec(x_40); -x_42 = l_Lean_IR_EmitLLVM_emitLhsSlotStore___redArg(x_1, x_2, x_3, x_41, x_6); -return x_42; +x_42 = lean_unbox_usize(x_41); +lean_dec(x_41); +x_43 = l_Lean_IR_EmitLLVM_emitLhsSlotStore___redArg(x_1, x_2, x_3, x_42, x_7); +return x_43; } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUProj___boxed(lean_object* x_1, lean_object* x_2, lean_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_emitUProj___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); +size_t x_10; size_t x_11; uint8_t x_12; lean_object* x_13; +x_10 = lean_unbox_usize(x_1); lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); +x_11 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = l_Lean_IR_EmitLLVM_emitUProj(x_9, x_10, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_7); -lean_dec(x_6); +x_12 = lean_unbox(x_5); +x_13 = l_Lean_IR_EmitLLVM_emitUProj(x_10, x_11, x_3, x_4, x_12, x_6, x_7, x_8); +lean_dec_ref(x_8); +lean_dec(x_7); lean_dec(x_4); -return x_11; +return x_13; } } LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitOffset___redArg(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { @@ -23024,7 +23130,7 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Invalid type for lean_ctor_get: '", 33, 33); +x_1 = lean_mk_string_unchecked("lean_ctor_get_int8", 18, 18); return x_1; } } @@ -23032,7 +23138,7 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__7() { _start: { lean_object* x_1; -x_1 = l_Std_Format_defWidth; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int16", 19, 19); return x_1; } } @@ -23040,6 +23146,38 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__8() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int32", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_get_int64", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Invalid type for lean_ctor_get: '", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Format_defWidth; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSProj___closed__12() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("'", 1, 1); return x_1; } @@ -23121,25 +23259,73 @@ x_14 = x_9; x_15 = lean_box(0); goto block_64; } +case 14: +{ +size_t x_77; lean_object* x_78; +x_77 = l_LLVM_i8Type(x_1); +x_78 = l_Lean_IR_EmitLLVM_emitSProj___closed__6; +x_11 = x_78; +x_12 = x_77; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_64; +} +case 15: +{ +size_t x_79; lean_object* x_80; +x_79 = l_LLVM_i16Type(x_1); +x_80 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; +x_11 = x_80; +x_12 = x_79; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_64; +} +case 16: +{ +size_t x_81; lean_object* x_82; +x_81 = l_LLVM_i32Type(x_1); +x_82 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; +x_11 = x_82; +x_12 = x_81; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_64; +} +case 17: +{ +size_t x_83; lean_object* x_84; +x_83 = l_LLVM_i64Type(x_1); +x_84 = l_Lean_IR_EmitLLVM_emitSProj___closed__9; +x_11 = x_84; +x_12 = x_83; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_64; +} default: { -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_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_dec(x_7); lean_dec(x_3); -x_77 = l_Lean_IR_EmitLLVM_emitSProj___closed__6; -x_78 = l_Lean_IR_instToFormatIRType___private__1(x_4); -x_79 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; -x_80 = lean_unsigned_to_nat(0u); -x_81 = l_Std_Format_pretty(x_78, x_79, x_80, x_80); -x_82 = lean_string_append(x_77, x_81); -lean_dec_ref(x_81); -x_83 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; -x_84 = lean_string_append(x_82, x_83); -x_85 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_85, 0, x_84); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -return x_86; +x_85 = l_Lean_IR_EmitLLVM_emitSProj___closed__10; +x_86 = l_Lean_IR_instToFormatIRType___private__1(x_4); +x_87 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; +x_88 = lean_unsigned_to_nat(0u); +x_89 = l_Std_Format_pretty(x_86, x_87, x_88, x_88); +x_90 = lean_string_append(x_85, x_89); +lean_dec_ref(x_89); +x_91 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; +x_92 = lean_string_append(x_90, x_91); +x_93 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_93, 0, x_92); +x_94 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_94, 0, x_93); +return x_94; } } block_64: @@ -23779,6 +23965,30 @@ x_1 = lean_mk_string_unchecked("lean_box_float32", 16, 16); return x_1; } } +static lean_object* _init_l_Lean_IR_EmitLLVM_emitBox___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_int32", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitBox___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_int64", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitBox___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_box_isize", 14, 14); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitBox(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: { @@ -23905,74 +24115,128 @@ x_14 = x_7; x_15 = lean_box(0); goto block_34; } -default: +case 16: { -lean_object* x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; lean_object* x_66; +lean_object* x_61; size_t x_62; lean_object* x_63; size_t x_64; x_61 = lean_ctor_get(x_37, 0); lean_inc(x_61); lean_dec_ref(x_37); -x_62 = l_LLVM_i64Type(x_1); -x_63 = lean_unbox_usize(x_61); +x_62 = l_LLVM_i32Type(x_1); +x_63 = l_Lean_IR_EmitLLVM_emitBox___closed__5; +x_64 = lean_unbox_usize(x_61); lean_dec(x_61); -x_64 = lean_llvm_build_sext(x_1, x_2, x_63, x_62, x_9); -x_65 = l_LLVM_i64Type(x_1); -x_66 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__0; -x_10 = x_66; -x_11 = x_65; +x_10 = x_63; +x_11 = x_62; x_12 = x_64; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); goto block_34; } +case 17: +{ +lean_object* x_65; size_t x_66; lean_object* x_67; size_t x_68; +x_65 = lean_ctor_get(x_37, 0); +lean_inc(x_65); +lean_dec_ref(x_37); +x_66 = l_LLVM_i64Type(x_1); +x_67 = l_Lean_IR_EmitLLVM_emitBox___closed__6; +x_68 = lean_unbox_usize(x_65); +lean_dec(x_65); +x_10 = x_67; +x_11 = x_66; +x_12 = x_68; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} +case 18: +{ +lean_object* x_69; size_t x_70; lean_object* x_71; size_t x_72; +x_69 = lean_ctor_get(x_37, 0); +lean_inc(x_69); +lean_dec_ref(x_37); +x_70 = l_LLVM_i64Type(x_1); +x_71 = l_Lean_IR_EmitLLVM_emitBox___closed__7; +x_72 = lean_unbox_usize(x_69); +lean_dec(x_69); +x_10 = x_71; +x_11 = x_70; +x_12 = x_72; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} +default: +{ +lean_object* x_73; size_t x_74; size_t x_75; size_t x_76; size_t x_77; lean_object* x_78; +x_73 = lean_ctor_get(x_37, 0); +lean_inc(x_73); +lean_dec_ref(x_37); +x_74 = l_LLVM_i64Type(x_1); +x_75 = lean_unbox_usize(x_73); +lean_dec(x_73); +x_76 = lean_llvm_build_sext(x_1, x_2, x_75, x_74, x_9); +x_77 = l_LLVM_i64Type(x_1); +x_78 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__0; +x_10 = x_78; +x_11 = x_77; +x_12 = x_76; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} } } } else { -lean_object* x_67; -x_67 = lean_ctor_get(x_35, 0); -lean_inc(x_67); +lean_object* x_79; +x_79 = lean_ctor_get(x_35, 0); +lean_inc(x_79); lean_dec(x_35); -if (lean_obj_tag(x_67) == 0) +if (lean_obj_tag(x_79) == 0) { -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; lean_object* x_83; lean_dec(x_3); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - x_69 = x_67; +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_67); - x_69 = lean_box(0); + lean_dec_ref(x_79); + x_81 = lean_box(0); } -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(0, 1, 0); } else { - x_70 = x_69; + x_82 = x_81; } -lean_ctor_set(x_70, 0, x_68); -x_71 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; +lean_ctor_set(x_82, 0, x_80); +x_83 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_83, 0, x_82); +return x_83; } else { switch (lean_obj_tag(x_5)) { case 0: { -lean_object* x_72; size_t x_73; lean_object* x_74; size_t x_75; -x_72 = lean_ctor_get(x_67, 0); -lean_inc(x_72); -lean_dec_ref(x_67); -x_73 = lean_llvm_double_type_in_context(x_1); -x_74 = l_Lean_IR_EmitLLVM_emitBox___closed__0; -x_75 = lean_unbox_usize(x_72); -lean_dec(x_72); -x_10 = x_74; -x_11 = x_73; -x_12 = x_75; +lean_object* x_84; size_t x_85; lean_object* x_86; size_t x_87; +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +lean_dec_ref(x_79); +x_85 = lean_llvm_double_type_in_context(x_1); +x_86 = l_Lean_IR_EmitLLVM_emitBox___closed__0; +x_87 = lean_unbox_usize(x_84); +lean_dec(x_84); +x_10 = x_86; +x_11 = x_85; +x_12 = x_87; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -23980,17 +24244,17 @@ goto block_34; } case 3: { -lean_object* x_76; size_t x_77; lean_object* x_78; size_t x_79; -x_76 = lean_ctor_get(x_67, 0); -lean_inc(x_76); -lean_dec_ref(x_67); -x_77 = l_LLVM_i32Type(x_1); -x_78 = l_Lean_IR_EmitLLVM_emitBox___closed__1; -x_79 = lean_unbox_usize(x_76); -lean_dec(x_76); -x_10 = x_78; -x_11 = x_77; -x_12 = x_79; +lean_object* x_88; size_t x_89; lean_object* x_90; size_t x_91; +x_88 = lean_ctor_get(x_79, 0); +lean_inc(x_88); +lean_dec_ref(x_79); +x_89 = l_LLVM_i32Type(x_1); +x_90 = l_Lean_IR_EmitLLVM_emitBox___closed__1; +x_91 = lean_unbox_usize(x_88); +lean_dec(x_88); +x_10 = x_90; +x_11 = x_89; +x_12 = x_91; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -23998,17 +24262,17 @@ goto block_34; } case 4: { -lean_object* x_80; size_t x_81; lean_object* x_82; size_t x_83; -x_80 = lean_ctor_get(x_67, 0); -lean_inc(x_80); -lean_dec_ref(x_67); -x_81 = l_LLVM_i64Type(x_1); -x_82 = l_Lean_IR_EmitLLVM_emitBox___closed__2; -x_83 = lean_unbox_usize(x_80); -lean_dec(x_80); -x_10 = x_82; -x_11 = x_81; -x_12 = x_83; +lean_object* x_92; size_t x_93; lean_object* x_94; size_t x_95; +x_92 = lean_ctor_get(x_79, 0); +lean_inc(x_92); +lean_dec_ref(x_79); +x_93 = l_LLVM_i64Type(x_1); +x_94 = l_Lean_IR_EmitLLVM_emitBox___closed__2; +x_95 = lean_unbox_usize(x_92); +lean_dec(x_92); +x_10 = x_94; +x_11 = x_93; +x_12 = x_95; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -24016,17 +24280,17 @@ goto block_34; } case 5: { -lean_object* x_84; size_t x_85; lean_object* x_86; size_t x_87; -x_84 = lean_ctor_get(x_67, 0); -lean_inc(x_84); -lean_dec_ref(x_67); -x_85 = l_LLVM_i64Type(x_1); -x_86 = l_Lean_IR_EmitLLVM_emitBox___closed__3; -x_87 = lean_unbox_usize(x_84); -lean_dec(x_84); -x_10 = x_86; -x_11 = x_85; -x_12 = x_87; +lean_object* x_96; size_t x_97; lean_object* x_98; size_t x_99; +x_96 = lean_ctor_get(x_79, 0); +lean_inc(x_96); +lean_dec_ref(x_79); +x_97 = l_LLVM_i64Type(x_1); +x_98 = l_Lean_IR_EmitLLVM_emitBox___closed__3; +x_99 = lean_unbox_usize(x_96); +lean_dec(x_96); +x_10 = x_98; +x_11 = x_97; +x_12 = x_99; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -24034,17 +24298,71 @@ goto block_34; } case 9: { -lean_object* x_88; size_t x_89; lean_object* x_90; size_t x_91; -x_88 = lean_ctor_get(x_67, 0); -lean_inc(x_88); -lean_dec_ref(x_67); -x_89 = lean_llvm_float_type_in_context(x_1); -x_90 = l_Lean_IR_EmitLLVM_emitBox___closed__4; -x_91 = lean_unbox_usize(x_88); -lean_dec(x_88); -x_10 = x_90; -x_11 = x_89; -x_12 = x_91; +lean_object* x_100; size_t x_101; lean_object* x_102; size_t x_103; +x_100 = lean_ctor_get(x_79, 0); +lean_inc(x_100); +lean_dec_ref(x_79); +x_101 = lean_llvm_float_type_in_context(x_1); +x_102 = l_Lean_IR_EmitLLVM_emitBox___closed__4; +x_103 = lean_unbox_usize(x_100); +lean_dec(x_100); +x_10 = x_102; +x_11 = x_101; +x_12 = x_103; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} +case 16: +{ +lean_object* x_104; size_t x_105; lean_object* x_106; size_t x_107; +x_104 = lean_ctor_get(x_79, 0); +lean_inc(x_104); +lean_dec_ref(x_79); +x_105 = l_LLVM_i32Type(x_1); +x_106 = l_Lean_IR_EmitLLVM_emitBox___closed__5; +x_107 = lean_unbox_usize(x_104); +lean_dec(x_104); +x_10 = x_106; +x_11 = x_105; +x_12 = x_107; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} +case 17: +{ +lean_object* x_108; size_t x_109; lean_object* x_110; size_t x_111; +x_108 = lean_ctor_get(x_79, 0); +lean_inc(x_108); +lean_dec_ref(x_79); +x_109 = l_LLVM_i64Type(x_1); +x_110 = l_Lean_IR_EmitLLVM_emitBox___closed__6; +x_111 = lean_unbox_usize(x_108); +lean_dec(x_108); +x_10 = x_110; +x_11 = x_109; +x_12 = x_111; +x_13 = x_6; +x_14 = x_7; +x_15 = lean_box(0); +goto block_34; +} +case 18: +{ +lean_object* x_112; size_t x_113; lean_object* x_114; size_t x_115; +x_112 = lean_ctor_get(x_79, 0); +lean_inc(x_112); +lean_dec_ref(x_79); +x_113 = l_LLVM_i64Type(x_1); +x_114 = l_Lean_IR_EmitLLVM_emitBox___closed__7; +x_115 = lean_unbox_usize(x_112); +lean_dec(x_112); +x_10 = x_114; +x_11 = x_113; +x_12 = x_115; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -24052,19 +24370,19 @@ goto block_34; } default: { -lean_object* x_92; size_t x_93; size_t x_94; size_t x_95; size_t x_96; lean_object* x_97; -x_92 = lean_ctor_get(x_67, 0); -lean_inc(x_92); -lean_dec_ref(x_67); -x_93 = l_LLVM_i64Type(x_1); -x_94 = lean_unbox_usize(x_92); -lean_dec(x_92); -x_95 = lean_llvm_build_sext(x_1, x_2, x_94, x_93, x_9); -x_96 = l_LLVM_i64Type(x_1); -x_97 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__0; -x_10 = x_97; -x_11 = x_96; -x_12 = x_95; +lean_object* x_116; size_t x_117; size_t x_118; size_t x_119; size_t x_120; lean_object* x_121; +x_116 = lean_ctor_get(x_79, 0); +lean_inc(x_116); +lean_dec_ref(x_79); +x_117 = l_LLVM_i64Type(x_1); +x_118 = lean_unbox_usize(x_116); +lean_dec(x_116); +x_119 = lean_llvm_build_sext(x_1, x_2, x_118, x_117, x_9); +x_120 = l_LLVM_i64Type(x_1); +x_121 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__0; +x_10 = x_121; +x_11 = x_120; +x_12 = x_119; x_13 = x_6; x_14 = x_7; x_15 = lean_box(0); @@ -24223,6 +24541,30 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__5() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_int32", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_int64", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_unbox_isize", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__8() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("lean_unbox", 10, 10); return x_1; } @@ -24837,14 +25179,377 @@ return x_154; } } } +case 16: +{ +lean_object* x_155; +lean_inc_ref(x_7); +x_155 = l_Lean_IR_EmitLLVM_toLLVMType(x_1, x_3, x_6, x_7); +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; +x_157 = lean_ctor_get(x_155, 0); +if (lean_obj_tag(x_157) == 0) +{ +uint8_t x_158; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_158 = !lean_is_exclusive(x_157); +if (x_158 == 0) +{ +return x_155; +} +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(0, 1, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_155, 0, x_160); +return x_155; +} +} +else +{ +lean_object* x_161; lean_object* x_162; size_t x_163; +lean_free_object(x_155); +x_161 = lean_ctor_get(x_157, 0); +lean_inc(x_161); +lean_dec_ref(x_157); +x_162 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__5; +x_163 = lean_unbox_usize(x_161); +lean_dec(x_161); +x_9 = x_162; +x_10 = x_163; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +else +{ +lean_object* x_164; +x_164 = lean_ctor_get(x_155, 0); +lean_inc(x_164); +lean_dec(x_155); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + x_166 = x_164; +} else { + lean_dec_ref(x_164); + x_166 = lean_box(0); +} +if (lean_is_scalar(x_166)) { + x_167 = lean_alloc_ctor(0, 1, 0); +} else { + x_167 = x_166; +} +lean_ctor_set(x_167, 0, x_165); +x_168 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_168, 0, x_167); +return x_168; +} +else +{ +lean_object* x_169; lean_object* x_170; size_t x_171; +x_169 = lean_ctor_get(x_164, 0); +lean_inc(x_169); +lean_dec_ref(x_164); +x_170 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__5; +x_171 = lean_unbox_usize(x_169); +lean_dec(x_169); +x_9 = x_170; +x_10 = x_171; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +} +else +{ +uint8_t x_172; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_172 = !lean_is_exclusive(x_155); +if (x_172 == 0) +{ +return x_155; +} +else +{ +lean_object* x_173; lean_object* x_174; +x_173 = lean_ctor_get(x_155, 0); +lean_inc(x_173); +lean_dec(x_155); +x_174 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_174, 0, x_173); +return x_174; +} +} +} +case 17: +{ +lean_object* x_175; +lean_inc_ref(x_7); +x_175 = l_Lean_IR_EmitLLVM_toLLVMType(x_1, x_3, x_6, x_7); +if (lean_obj_tag(x_175) == 0) +{ +uint8_t x_176; +x_176 = !lean_is_exclusive(x_175); +if (x_176 == 0) +{ +lean_object* x_177; +x_177 = lean_ctor_get(x_175, 0); +if (lean_obj_tag(x_177) == 0) +{ +uint8_t x_178; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_178 = !lean_is_exclusive(x_177); +if (x_178 == 0) +{ +return x_175; +} +else +{ +lean_object* x_179; lean_object* x_180; +x_179 = lean_ctor_get(x_177, 0); +lean_inc(x_179); +lean_dec(x_177); +x_180 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_175, 0, x_180); +return x_175; +} +} +else +{ +lean_object* x_181; lean_object* x_182; size_t x_183; +lean_free_object(x_175); +x_181 = lean_ctor_get(x_177, 0); +lean_inc(x_181); +lean_dec_ref(x_177); +x_182 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__6; +x_183 = lean_unbox_usize(x_181); +lean_dec(x_181); +x_9 = x_182; +x_10 = x_183; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +else +{ +lean_object* x_184; +x_184 = lean_ctor_get(x_175, 0); +lean_inc(x_184); +lean_dec(x_175); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(0, 1, 0); +} else { + x_187 = x_186; +} +lean_ctor_set(x_187, 0, x_185); +x_188 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_188, 0, x_187); +return x_188; +} +else +{ +lean_object* x_189; lean_object* x_190; size_t x_191; +x_189 = lean_ctor_get(x_184, 0); +lean_inc(x_189); +lean_dec_ref(x_184); +x_190 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__6; +x_191 = lean_unbox_usize(x_189); +lean_dec(x_189); +x_9 = x_190; +x_10 = x_191; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +} +else +{ +uint8_t x_192; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_192 = !lean_is_exclusive(x_175); +if (x_192 == 0) +{ +return x_175; +} +else +{ +lean_object* x_193; lean_object* x_194; +x_193 = lean_ctor_get(x_175, 0); +lean_inc(x_193); +lean_dec(x_175); +x_194 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_194, 0, x_193); +return x_194; +} +} +} +case 18: +{ +lean_object* x_195; +lean_inc_ref(x_7); +x_195 = l_Lean_IR_EmitLLVM_toLLVMType(x_1, x_3, x_6, x_7); +if (lean_obj_tag(x_195) == 0) +{ +uint8_t x_196; +x_196 = !lean_is_exclusive(x_195); +if (x_196 == 0) +{ +lean_object* x_197; +x_197 = lean_ctor_get(x_195, 0); +if (lean_obj_tag(x_197) == 0) +{ +uint8_t x_198; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_198 = !lean_is_exclusive(x_197); +if (x_198 == 0) +{ +return x_195; +} +else +{ +lean_object* x_199; lean_object* x_200; +x_199 = lean_ctor_get(x_197, 0); +lean_inc(x_199); +lean_dec(x_197); +x_200 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_195, 0, x_200); +return x_195; +} +} +else +{ +lean_object* x_201; lean_object* x_202; size_t x_203; +lean_free_object(x_195); +x_201 = lean_ctor_get(x_197, 0); +lean_inc(x_201); +lean_dec_ref(x_197); +x_202 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__7; +x_203 = lean_unbox_usize(x_201); +lean_dec(x_201); +x_9 = x_202; +x_10 = x_203; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +else +{ +lean_object* x_204; +x_204 = lean_ctor_get(x_195, 0); +lean_inc(x_204); +lean_dec(x_195); +if (lean_obj_tag(x_204) == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_205 = lean_ctor_get(x_204, 0); +lean_inc(x_205); +if (lean_is_exclusive(x_204)) { + lean_ctor_release(x_204, 0); + x_206 = x_204; +} else { + lean_dec_ref(x_204); + x_206 = lean_box(0); +} +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(0, 1, 0); +} else { + x_207 = x_206; +} +lean_ctor_set(x_207, 0, x_205); +x_208 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_208, 0, x_207); +return x_208; +} +else +{ +lean_object* x_209; lean_object* x_210; size_t x_211; +x_209 = lean_ctor_get(x_204, 0); +lean_inc(x_209); +lean_dec_ref(x_204); +x_210 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__7; +x_211 = lean_unbox_usize(x_209); +lean_dec(x_209); +x_9 = x_210; +x_10 = x_211; +x_11 = x_7; +x_12 = lean_box(0); +goto block_54; +} +} +} +else +{ +uint8_t x_212; +lean_dec_ref(x_7); +lean_dec_ref(x_5); +x_212 = !lean_is_exclusive(x_195); +if (x_212 == 0) +{ +return x_195; +} +else +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_ctor_get(x_195, 0); +lean_inc(x_213); +lean_dec(x_195); +x_214 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_214, 0, x_213); +return x_214; +} +} +} default: { -size_t x_155; lean_object* x_156; +size_t x_215; lean_object* x_216; lean_dec(x_6); -x_155 = l_LLVM_i64Type(x_1); -x_156 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__5; -x_9 = x_156; -x_10 = x_155; +x_215 = l_LLVM_i64Type(x_1); +x_216 = l_Lean_IR_EmitLLVM_callUnboxForType___closed__8; +x_9 = x_216; +x_10 = x_215; x_11 = x_7; x_12 = lean_box(0); goto block_54; @@ -27277,216 +27982,217 @@ return x_22; } case 4: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_dec(x_4); x_23 = lean_ctor_get(x_5, 0); lean_inc(x_23); -x_24 = lean_ctor_get(x_5, 1); -lean_inc(x_24); +x_24 = lean_ctor_get_uint8(x_5, sizeof(void*)*2); +x_25 = lean_ctor_get(x_5, 1); +lean_inc(x_25); 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); +x_26 = l_Lean_IR_EmitLLVM_emitUProj(x_1, x_2, x_3, x_23, x_24, x_25, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); lean_dec(x_23); -return x_25; +return x_26; } 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); -lean_inc(x_26); -x_27 = lean_ctor_get(x_5, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_5, 0); lean_inc(x_27); -x_28 = lean_ctor_get(x_5, 2); +x_28 = lean_ctor_get(x_5, 1); lean_inc(x_28); +x_29 = lean_ctor_get(x_5, 2); +lean_inc(x_29); 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); +x_30 = l_Lean_IR_EmitLLVM_emitSProj(x_1, x_2, x_3, x_4, x_27, x_28, x_29, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); +lean_dec(x_28); lean_dec(x_27); -lean_dec(x_26); -return x_29; +return x_30; } case 6: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_4); -x_30 = lean_ctor_get(x_5, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_5, 1); -lean_inc_ref(x_31); +x_31 = lean_ctor_get(x_5, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_5, 1); +lean_inc_ref(x_32); lean_dec_ref(x_5); -x_32 = l_Lean_IR_EmitLLVM_emitFullApp(x_1, x_2, x_3, x_30, x_31, x_6, x_7); -return x_32; +x_33 = l_Lean_IR_EmitLLVM_emitFullApp(x_1, x_2, x_3, x_31, x_32, x_6, x_7); +return x_33; } case 7: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_4); -x_33 = lean_ctor_get(x_5, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_5, 1); -lean_inc_ref(x_34); +x_34 = lean_ctor_get(x_5, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_5, 1); +lean_inc_ref(x_35); lean_dec_ref(x_5); -x_35 = l_Lean_IR_EmitLLVM_emitPartialApp(x_1, x_2, x_3, x_33, x_34, x_6, x_7); -lean_dec_ref(x_34); -return x_35; +x_36 = l_Lean_IR_EmitLLVM_emitPartialApp(x_1, x_2, x_3, x_34, x_35, x_6, x_7); +lean_dec_ref(x_35); +return x_36; } case 8: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_dec(x_4); -x_36 = lean_ctor_get(x_5, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_5, 1); -lean_inc_ref(x_37); +x_37 = lean_ctor_get(x_5, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_5, 1); +lean_inc_ref(x_38); lean_dec_ref(x_5); -x_38 = l_Lean_IR_EmitLLVM_emitApp(x_1, x_2, x_3, x_36, x_37, x_6, x_7); +x_39 = l_Lean_IR_EmitLLVM_emitApp(x_1, x_2, x_3, x_37, x_38, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); -return x_38; +return x_39; } case 9: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_dec(x_4); -x_39 = lean_ctor_get(x_5, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_5, 1); +x_40 = lean_ctor_get(x_5, 0); lean_inc(x_40); +x_41 = lean_ctor_get(x_5, 1); +lean_inc(x_41); lean_dec_ref(x_5); -x_41 = l_Lean_IR_EmitLLVM_emitBox(x_1, x_2, x_3, x_40, x_39, x_6, x_7); +x_42 = l_Lean_IR_EmitLLVM_emitBox(x_1, x_2, x_3, x_41, x_40, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); -lean_dec(x_39); -return x_41; +lean_dec(x_40); +return x_42; } case 10: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_5, 0); -lean_inc(x_42); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_5, 0); +lean_inc(x_43); lean_dec_ref(x_5); -x_43 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; -x_44 = l_Lean_IR_EmitLLVM_emitUnbox(x_1, x_2, x_3, x_4, x_42, x_43, x_6, x_7); +x_44 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; +x_45 = l_Lean_IR_EmitLLVM_emitUnbox(x_1, x_2, x_3, x_4, x_43, x_44, x_6, x_7); lean_dec(x_4); -return x_44; +return x_45; } case 11: { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_5, 0); -lean_inc_ref(x_45); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_5, 0); +lean_inc_ref(x_46); lean_dec_ref(x_5); -x_46 = l_Lean_IR_EmitLLVM_emitLit(x_1, x_2, x_3, x_4, x_45, x_6, x_7); +x_47 = l_Lean_IR_EmitLLVM_emitLit(x_1, x_2, x_3, x_4, x_46, x_6, x_7); lean_dec(x_4); -if (lean_obj_tag(x_46) == 0) +if (lean_obj_tag(x_47) == 0) { -uint8_t x_47; -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_46, 0); -if (lean_obj_tag(x_48) == 0) +lean_object* x_49; +x_49 = lean_ctor_get(x_47, 0); +if (lean_obj_tag(x_49) == 0) { -uint8_t x_49; -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) { -return x_46; +return x_47; } 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(0, 1, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_46, 0, x_51); -return x_46; +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_47, 0, x_52); +return x_47; } } else { -lean_object* x_52; -lean_dec_ref(x_48); -x_52 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -lean_ctor_set(x_46, 0, x_52); -return x_46; +lean_object* x_53; +lean_dec_ref(x_49); +x_53 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +lean_ctor_set(x_47, 0, x_53); +return x_47; } } else { -lean_object* x_53; -x_53 = lean_ctor_get(x_46, 0); -lean_inc(x_53); -lean_dec(x_46); -if (lean_obj_tag(x_53) == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_53, 0); +lean_object* x_54; +x_54 = lean_ctor_get(x_47, 0); lean_inc(x_54); -if (lean_is_exclusive(x_53)) { - lean_ctor_release(x_53, 0); - x_55 = x_53; +lean_dec(x_47); +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); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + x_56 = x_54; } else { - lean_dec_ref(x_53); - x_55 = lean_box(0); + lean_dec_ref(x_54); + x_56 = lean_box(0); } -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 1, 0); } else { - x_56 = x_55; + x_57 = x_56; } -lean_ctor_set(x_56, 0, x_54); -x_57 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_57, 0, x_56); -return x_57; +lean_ctor_set(x_57, 0, x_55); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_57); +return x_58; } else { -lean_object* x_58; lean_object* x_59; -lean_dec_ref(x_53); -x_58 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -x_59 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_59, 0, x_58); -return x_59; +lean_object* x_59; lean_object* x_60; +lean_dec_ref(x_54); +x_59 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_60, 0, x_59); +return x_60; } } } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_46); -if (x_60 == 0) +uint8_t x_61; +x_61 = !lean_is_exclusive(x_47); +if (x_61 == 0) { -return x_46; +return x_47; } else { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_46, 0); -lean_inc(x_61); -lean_dec(x_46); -x_62 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_47, 0); +lean_inc(x_62); +lean_dec(x_47); +x_63 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_63, 0, x_62); +return x_63; } } } default: { -lean_object* x_63; lean_object* x_64; +lean_object* x_64; lean_object* x_65; lean_dec(x_4); -x_63 = lean_ctor_get(x_5, 0); -lean_inc(x_63); +x_64 = lean_ctor_get(x_5, 0); +lean_inc(x_64); lean_dec_ref(x_5); -x_64 = l_Lean_IR_EmitLLVM_emitIsShared(x_1, x_2, x_3, x_63, x_6, x_7); +x_65 = l_Lean_IR_EmitLLVM_emitIsShared(x_1, x_2, x_3, x_64, x_6, x_7); lean_dec_ref(x_7); lean_dec(x_6); -return x_64; +return x_65; } } } @@ -28396,304 +29102,331 @@ x_1 = lean_mk_string_unchecked("lean_ctor_set_usize", 19, 19); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet(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) { +static lean_object* _init_l_Lean_IR_EmitLLVM_emitUSet___closed__1() { _start: { -size_t x_9; size_t x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_9 = lean_llvm_void_type_in_context(x_1); -x_10 = l_LLVM_voidPtrType(x_1); -x_11 = l_LLVM_i32Type(x_1); -x_12 = l_LLVM_i64Type(x_1); -x_13 = l_Lean_IR_EmitLLVM_getLLVMModule___redArg(x_7); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec_ref(x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -x_16 = l_Lean_IR_EmitLLVM_emitUSet___closed__0; -x_17 = l_Lean_IR_EmitLLVM_callLeanMkStringUncheckedFn___redArg___closed__1; -x_18 = lean_box_usize(x_10); -x_19 = lean_array_push(x_17, x_18); -x_20 = lean_box_usize(x_11); -x_21 = lean_array_push(x_19, x_20); -x_22 = lean_box_usize(x_12); -x_23 = lean_array_push(x_21, x_22); -x_24 = lean_unbox_usize(x_15); -lean_dec(x_15); -x_25 = l_Lean_IR_EmitLLVM_getOrCreateFunctionPrototype___redArg(x_1, x_24, x_9, x_16, x_23); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec_ref(x_25); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_isize", 19, 19); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet(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, lean_object* x_8) { +_start: +{ +lean_object* x_10; +if (x_5 == 0) +{ +lean_object* x_93; +x_93 = l_Lean_IR_EmitLLVM_emitUSet___closed__0; +x_10 = x_93; +goto block_92; +} +else +{ +lean_object* x_94; +x_94 = l_Lean_IR_EmitLLVM_emitUSet___closed__1; +x_10 = x_94; +goto block_92; +} +block_92: +{ +size_t x_11; size_t x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_11 = lean_llvm_void_type_in_context(x_1); +x_12 = l_LLVM_voidPtrType(x_1); +x_13 = l_LLVM_i32Type(x_1); +x_14 = l_LLVM_i64Type(x_1); +x_15 = l_Lean_IR_EmitLLVM_getLLVMModule___redArg(x_8); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec_ref(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_IR_EmitLLVM_callLeanMkStringUncheckedFn___redArg___closed__1; +x_19 = lean_box_usize(x_12); +x_20 = lean_array_push(x_18, x_19); +x_21 = lean_box_usize(x_13); +x_22 = lean_array_push(x_20, x_21); +x_23 = lean_box_usize(x_14); +x_24 = lean_array_push(x_22, x_23); +x_25 = lean_unbox_usize(x_17); +lean_dec(x_17); +x_26 = l_Lean_IR_EmitLLVM_getOrCreateFunctionPrototype___redArg(x_1, x_25, x_11, x_10, x_24); +lean_dec_ref(x_10); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_26); -x_28 = 0; -x_29 = lean_llvm_function_type(x_1, x_9, x_23, x_28); -lean_dec_ref(x_23); -x_30 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; -x_31 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_3, x_30, x_6); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_dec_ref(x_26); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +x_29 = 0; +x_30 = lean_llvm_function_type(x_1, x_11, x_24, x_29); +lean_dec_ref(x_24); +x_31 = l_Lean_IR_EmitLLVM_callLeanMarkPersistentFn___redArg___closed__1; +x_32 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_3, x_31, x_7); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_31, 0); -if (lean_obj_tag(x_33) == 0) +lean_object* x_34; +x_34 = lean_ctor_get(x_32, 0); +if (lean_obj_tag(x_34) == 0) { -uint8_t x_34; -lean_dec(x_27); -lean_dec(x_5); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +uint8_t x_35; +lean_dec(x_28); +lean_dec(x_6); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) { -return x_31; +return x_32; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_31, 0, x_36); -return x_31; +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_32, 0, x_37); +return x_32; } } else { -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_free_object(x_31); -x_37 = lean_ctor_get(x_33, 0); -lean_inc(x_37); -lean_dec_ref(x_33); -x_38 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -lean_dec_ref(x_38); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +lean_free_object(x_32); +x_38 = lean_ctor_get(x_34, 0); +lean_inc(x_38); +lean_dec_ref(x_34); +x_39 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_39); -x_41 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_5, x_30, x_6); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +lean_dec_ref(x_39); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_6, x_31, x_7); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_43; -x_43 = lean_ctor_get(x_41, 0); -if (lean_obj_tag(x_43) == 0) +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +if (lean_obj_tag(x_44) == 0) { -uint8_t x_44; -lean_dec(x_40); -lean_dec(x_37); -lean_dec(x_27); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_45; +lean_dec(x_41); +lean_dec(x_38); +lean_dec(x_28); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) { -return x_41; +return x_42; } else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_43, 0); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_41, 0, x_46); -return x_41; +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_42, 0, x_47); +return x_42; } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; lean_object* x_53; -x_47 = lean_ctor_get(x_43, 0); -lean_inc(x_47); -lean_dec_ref(x_43); -x_48 = lean_array_push(x_17, x_37); -x_49 = lean_array_push(x_48, x_40); -x_50 = lean_array_push(x_49, x_47); -x_51 = lean_unbox_usize(x_27); -lean_dec(x_27); -x_52 = lean_llvm_build_call2(x_1, x_2, x_29, x_51, x_50, x_30); -lean_dec_ref(x_50); -x_53 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -lean_ctor_set(x_41, 0, x_53); -return x_41; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; +x_48 = lean_ctor_get(x_44, 0); +lean_inc(x_48); +lean_dec_ref(x_44); +x_49 = lean_array_push(x_18, x_38); +x_50 = lean_array_push(x_49, x_41); +x_51 = lean_array_push(x_50, x_48); +x_52 = lean_unbox_usize(x_28); +lean_dec(x_28); +x_53 = lean_llvm_build_call2(x_1, x_2, x_30, x_52, x_51, x_31); +lean_dec_ref(x_51); +x_54 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +lean_ctor_set(x_42, 0, x_54); +return x_42; } } else { -lean_object* x_54; -x_54 = lean_ctor_get(x_41, 0); -lean_inc(x_54); -lean_dec(x_41); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_40); -lean_dec(x_37); -lean_dec(x_27); -x_55 = lean_ctor_get(x_54, 0); +lean_object* x_55; +x_55 = lean_ctor_get(x_42, 0); lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - x_56 = x_54; +lean_dec(x_42); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_41); +lean_dec(x_38); +lean_dec(x_28); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + x_57 = x_55; } else { - lean_dec_ref(x_54); - x_56 = lean_box(0); + lean_dec_ref(x_55); + x_57 = lean_box(0); } -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(0, 1, 0); } else { - x_57 = x_56; + x_58 = x_57; } -lean_ctor_set(x_57, 0, x_55); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_57); -return x_58; +lean_ctor_set(x_58, 0, x_56); +x_59 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_59, 0, x_58); +return x_59; } else { -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; -x_59 = lean_ctor_get(x_54, 0); -lean_inc(x_59); -lean_dec_ref(x_54); -x_60 = lean_array_push(x_17, x_37); -x_61 = lean_array_push(x_60, x_40); -x_62 = lean_array_push(x_61, x_59); -x_63 = lean_unbox_usize(x_27); -lean_dec(x_27); -x_64 = lean_llvm_build_call2(x_1, x_2, x_29, x_63, x_62, x_30); -lean_dec_ref(x_62); -x_65 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -x_66 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_66, 0, x_65); -return x_66; +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; +x_60 = lean_ctor_get(x_55, 0); +lean_inc(x_60); +lean_dec_ref(x_55); +x_61 = lean_array_push(x_18, x_38); +x_62 = lean_array_push(x_61, x_41); +x_63 = lean_array_push(x_62, x_60); +x_64 = lean_unbox_usize(x_28); +lean_dec(x_28); +x_65 = lean_llvm_build_call2(x_1, x_2, x_30, x_64, x_63, x_31); +lean_dec_ref(x_63); +x_66 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } } } } else { -lean_object* x_67; -x_67 = lean_ctor_get(x_31, 0); -lean_inc(x_67); -lean_dec(x_31); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -lean_dec(x_27); -lean_dec(x_5); -x_68 = lean_ctor_get(x_67, 0); +lean_object* x_68; +x_68 = lean_ctor_get(x_32, 0); lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - x_69 = x_67; +lean_dec(x_32); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_28); +lean_dec(x_6); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + x_70 = x_68; } else { - lean_dec_ref(x_67); - x_69 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(0, 1, 0); } else { - x_70 = x_69; + x_71 = x_70; } -lean_ctor_set(x_70, 0, x_68); -x_71 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_71, 0, x_70); -return x_71; +lean_ctor_set(x_71, 0, x_69); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_71); +return x_72; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_72 = lean_ctor_get(x_67, 0); -lean_inc(x_72); -lean_dec_ref(x_67); -x_73 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -lean_dec_ref(x_73); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_73 = lean_ctor_get(x_68, 0); +lean_inc(x_73); +lean_dec_ref(x_68); +x_74 = l_Lean_IR_EmitLLVM_constIntUnsigned___redArg(x_1, x_4); x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); -lean_dec(x_74); -x_76 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_5, x_30, x_6); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - x_78 = x_76; -} else { - lean_dec_ref(x_76); - x_78 = lean_box(0); -} -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec_ref(x_74); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); lean_dec(x_75); -lean_dec(x_72); -lean_dec(x_27); -x_79 = lean_ctor_get(x_77, 0); -lean_inc(x_79); +x_77 = l_Lean_IR_EmitLLVM_emitLhsVal___redArg(x_1, x_2, x_6, x_31, x_7); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); if (lean_is_exclusive(x_77)) { lean_ctor_release(x_77, 0); - x_80 = x_77; + x_79 = x_77; } else { lean_dec_ref(x_77); - x_80 = lean_box(0); + x_79 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(0, 1, 0); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_76); +lean_dec(x_73); +lean_dec(x_28); +x_80 = lean_ctor_get(x_78, 0); +lean_inc(x_80); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + x_81 = x_78; } else { - x_81 = x_80; + lean_dec_ref(x_78); + x_81 = lean_box(0); } -lean_ctor_set(x_81, 0, x_79); -if (lean_is_scalar(x_78)) { +if (lean_is_scalar(x_81)) { x_82 = lean_alloc_ctor(0, 1, 0); } else { - x_82 = x_78; + x_82 = x_81; } -lean_ctor_set(x_82, 0, x_81); -return x_82; +lean_ctor_set(x_82, 0, x_80); +if (lean_is_scalar(x_79)) { + x_83 = lean_alloc_ctor(0, 1, 0); +} else { + x_83 = x_79; +} +lean_ctor_set(x_83, 0, x_82); +return x_83; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; -x_83 = lean_ctor_get(x_77, 0); -lean_inc(x_83); -lean_dec_ref(x_77); -x_84 = lean_array_push(x_17, x_72); -x_85 = lean_array_push(x_84, x_75); -x_86 = lean_array_push(x_85, x_83); -x_87 = lean_unbox_usize(x_27); -lean_dec(x_27); -x_88 = lean_llvm_build_call2(x_1, x_2, x_29, x_87, x_86, x_30); -lean_dec_ref(x_86); -x_89 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -if (lean_is_scalar(x_78)) { - x_90 = lean_alloc_ctor(0, 1, 0); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; lean_object* x_90; lean_object* x_91; +x_84 = lean_ctor_get(x_78, 0); +lean_inc(x_84); +lean_dec_ref(x_78); +x_85 = lean_array_push(x_18, x_73); +x_86 = lean_array_push(x_85, x_76); +x_87 = lean_array_push(x_86, x_84); +x_88 = lean_unbox_usize(x_28); +lean_dec(x_28); +x_89 = lean_llvm_build_call2(x_1, x_2, x_30, x_88, x_87, x_31); +lean_dec_ref(x_87); +x_90 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +if (lean_is_scalar(x_79)) { + x_91 = lean_alloc_ctor(0, 1, 0); } else { - x_90 = x_78; + x_91 = x_79; +} +lean_ctor_set(x_91, 0, x_90); +return x_91; } -lean_ctor_set(x_90, 0, x_89); -return x_90; } } } } } -LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitUSet___boxed(lean_object* x_1, lean_object* x_2, lean_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_emitUSet___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); +size_t x_10; size_t x_11; uint8_t x_12; lean_object* x_13; +x_10 = lean_unbox_usize(x_1); lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); +x_11 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = l_Lean_IR_EmitLLVM_emitUSet(x_9, x_10, x_3, x_4, x_5, x_6, x_7); -lean_dec_ref(x_7); -lean_dec(x_6); +x_12 = lean_unbox(x_5); +x_13 = l_Lean_IR_EmitLLVM_emitUSet(x_10, x_11, x_3, x_4, x_12, x_6, x_7, x_8); +lean_dec_ref(x_8); +lean_dec(x_7); lean_dec(x_4); -return x_11; +return x_13; } } LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_EmitLLVM_emitTailCall_spec__0(size_t x_1, size_t x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { @@ -28863,12 +29596,12 @@ lean_dec_ref(x_15); x_17 = l_Lean_IR_EmitLLVM_emitTailCall___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr(x_4); -x_20 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; +x_20 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; x_21 = lean_unsigned_to_nat(0u); x_22 = l_Std_Format_pretty(x_19, x_20, x_21, x_21); x_23 = lean_string_append(x_18, x_22); lean_dec_ref(x_22); -x_24 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; +x_24 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; x_25 = lean_string_append(x_23, x_24); x_26 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_26, 0, x_25); @@ -29196,12 +29929,12 @@ lean_dec(x_5); lean_dec(x_3); x_99 = l_Lean_IR_EmitLLVM_emitTailCall___closed__2; x_100 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr(x_4); -x_101 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; +x_101 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; x_102 = lean_unsigned_to_nat(0u); x_103 = l_Std_Format_pretty(x_100, x_101, x_102, x_102); x_104 = lean_string_append(x_99, x_103); lean_dec_ref(x_103); -x_105 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; +x_105 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; x_106 = lean_string_append(x_104, x_105); x_107 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_107, 0, x_106); @@ -29706,6 +30439,38 @@ static lean_object* _init_l_Lean_IR_EmitLLVM_emitSSet___closed__6() { _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int8", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSSet___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int16", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSSet___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int32", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSSet___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lean_ctor_set_int64", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitLLVM_emitSSet___closed__10() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("invalid type for 'lean_ctor_set': '", 35, 35); return x_1; } @@ -29787,25 +30552,73 @@ x_14 = x_9; x_15 = lean_box(0); goto block_100; } +case 14: +{ +size_t x_113; lean_object* x_114; +x_113 = l_LLVM_i8Type(x_1); +x_114 = l_Lean_IR_EmitLLVM_emitSSet___closed__6; +x_11 = x_114; +x_12 = x_113; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_100; +} +case 15: +{ +size_t x_115; lean_object* x_116; +x_115 = l_LLVM_i16Type(x_1); +x_116 = l_Lean_IR_EmitLLVM_emitSSet___closed__7; +x_11 = x_116; +x_12 = x_115; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_100; +} +case 16: +{ +size_t x_117; lean_object* x_118; +x_117 = l_LLVM_i32Type(x_1); +x_118 = l_Lean_IR_EmitLLVM_emitSSet___closed__8; +x_11 = x_118; +x_12 = x_117; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_100; +} +case 17: +{ +size_t x_119; lean_object* x_120; +x_119 = l_LLVM_i64Type(x_1); +x_120 = l_Lean_IR_EmitLLVM_emitSSet___closed__9; +x_11 = x_120; +x_12 = x_119; +x_13 = x_8; +x_14 = x_9; +x_15 = lean_box(0); +goto block_100; +} default: { -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_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_6); lean_dec(x_3); -x_113 = l_Lean_IR_EmitLLVM_emitSSet___closed__6; -x_114 = l_Lean_IR_instToFormatIRType___private__1(x_7); -x_115 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; -x_116 = lean_unsigned_to_nat(0u); -x_117 = l_Std_Format_pretty(x_114, x_115, x_116, x_116); -x_118 = lean_string_append(x_113, x_117); -lean_dec_ref(x_117); -x_119 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; -x_120 = lean_string_append(x_118, x_119); -x_121 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_121, 0, x_120); -x_122 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_122, 0, x_121); -return x_122; +x_121 = l_Lean_IR_EmitLLVM_emitSSet___closed__10; +x_122 = l_Lean_IR_instToFormatIRType___private__1(x_7); +x_123 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; +x_124 = lean_unsigned_to_nat(0u); +x_125 = l_Std_Format_pretty(x_122, x_123, x_124, x_124); +x_126 = lean_string_append(x_121, x_125); +lean_dec_ref(x_125); +x_127 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; +x_128 = lean_string_append(x_126, x_127); +x_129 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_129, 0, x_128); +x_130 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_130, 0, x_129); +return x_130; } } block_100: @@ -30871,7 +31684,7 @@ lean_dec(x_25); x_29 = lean_llvm_build_zext(x_1, x_2, x_28, x_26, x_27); x_30 = l_Lean_IR_EmitLLVM_emitCase___closed__0; x_31 = l_Lean_IR_instToFormatIRType___private__1(x_4); -x_32 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; +x_32 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; x_33 = lean_unsigned_to_nat(0u); x_34 = l_Std_Format_pretty(x_31, x_32, x_33, x_33); x_35 = lean_string_append(x_30, x_34); @@ -31001,7 +31814,7 @@ lean_dec(x_59); x_63 = lean_llvm_build_zext(x_1, x_2, x_62, x_60, x_61); x_64 = l_Lean_IR_EmitLLVM_emitCase___closed__0; x_65 = l_Lean_IR_instToFormatIRType___private__1(x_4); -x_66 = l_Lean_IR_EmitLLVM_emitSProj___closed__7; +x_66 = l_Lean_IR_EmitLLVM_emitSProj___closed__11; x_67 = lean_unsigned_to_nat(0u); x_68 = l_Std_Format_pretty(x_65, x_66, x_67, x_67); x_69 = lean_string_append(x_64, x_68); @@ -31318,323 +32131,324 @@ goto _start; } case 4: { -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_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_36 = lean_ctor_get(x_3, 0); lean_inc(x_36); x_37 = lean_ctor_get(x_3, 1); lean_inc(x_37); x_38 = lean_ctor_get(x_3, 2); lean_inc(x_38); -x_39 = lean_ctor_get(x_3, 3); -lean_inc(x_39); +x_39 = lean_ctor_get_uint8(x_3, sizeof(void*)*4); +x_40 = lean_ctor_get(x_3, 3); +lean_inc(x_40); 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); +x_41 = l_Lean_IR_EmitLLVM_emitUSet(x_1, x_2, x_36, x_37, x_39, x_38, x_4, x_5); lean_dec(x_37); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +if (lean_obj_tag(x_42) == 0) { -lean_dec_ref(x_41); -lean_dec(x_39); +lean_dec_ref(x_42); +lean_dec(x_40); lean_dec_ref(x_5); lean_dec(x_4); -return x_40; +return x_41; } else { +lean_dec_ref(x_42); lean_dec_ref(x_41); -lean_dec_ref(x_40); -x_3 = x_39; +x_3 = x_40; goto _start; } } 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; lean_object* x_50; -x_43 = lean_ctor_get(x_3, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_3, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = lean_ctor_get(x_3, 0); lean_inc(x_44); -x_45 = lean_ctor_get(x_3, 2); +x_45 = lean_ctor_get(x_3, 1); lean_inc(x_45); -x_46 = lean_ctor_get(x_3, 3); +x_46 = lean_ctor_get(x_3, 2); lean_inc(x_46); -x_47 = lean_ctor_get(x_3, 4); +x_47 = lean_ctor_get(x_3, 3); lean_inc(x_47); -x_48 = lean_ctor_get(x_3, 5); +x_48 = lean_ctor_get(x_3, 4); lean_inc(x_48); +x_49 = lean_ctor_get(x_3, 5); +lean_inc(x_49); 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); +x_50 = l_Lean_IR_EmitLLVM_emitSSet(x_1, x_2, x_44, x_45, x_46, x_47, x_48, x_4, x_5); +lean_dec(x_46); lean_dec(x_45); -lean_dec(x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -if (lean_obj_tag(x_50) == 0) +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 0) { -lean_dec_ref(x_50); -lean_dec(x_48); +lean_dec_ref(x_51); +lean_dec(x_49); lean_dec_ref(x_5); lean_dec(x_4); -return x_49; +return x_50; } else { +lean_dec_ref(x_51); lean_dec_ref(x_50); -lean_dec_ref(x_49); -x_3 = x_48; +x_3 = x_49; goto _start; } } case 6: { -uint8_t x_52; -x_52 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); -if (x_52 == 0) +uint8_t x_53; +x_53 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); +if (x_53 == 0) { -lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_53 = lean_ctor_get(x_3, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_3, 1); +lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_54 = lean_ctor_get(x_3, 0); lean_inc(x_54); -x_55 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_56 = lean_ctor_get(x_3, 2); -lean_inc(x_56); +x_55 = lean_ctor_get(x_3, 1); +lean_inc(x_55); +x_56 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_57 = lean_ctor_get(x_3, 2); +lean_inc(x_57); lean_dec_ref(x_3); -x_57 = l_Lean_IR_EmitLLVM_emitInc(x_1, x_2, x_53, x_54, x_55, x_4, x_5); -lean_dec(x_54); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -if (lean_obj_tag(x_58) == 0) +x_58 = l_Lean_IR_EmitLLVM_emitInc(x_1, x_2, x_54, x_55, x_56, x_4, x_5); +lean_dec(x_55); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +if (lean_obj_tag(x_59) == 0) { -lean_dec_ref(x_58); -lean_dec(x_56); +lean_dec_ref(x_59); +lean_dec(x_57); lean_dec_ref(x_5); lean_dec(x_4); -return x_57; +return x_58; } else { +lean_dec_ref(x_59); lean_dec_ref(x_58); -lean_dec_ref(x_57); -x_3 = x_56; +x_3 = x_57; goto _start; } } else { -lean_object* x_60; -x_60 = lean_ctor_get(x_3, 2); -lean_inc(x_60); +lean_object* x_61; +x_61 = lean_ctor_get(x_3, 2); +lean_inc(x_61); lean_dec_ref(x_3); -x_3 = x_60; +x_3 = x_61; goto _start; } } case 7: { -uint8_t x_62; -x_62 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); -if (x_62 == 0) +uint8_t x_63; +x_63 = lean_ctor_get_uint8(x_3, sizeof(void*)*3 + 1); +if (x_63 == 0) { -lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_63 = lean_ctor_get(x_3, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_3, 1); +lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_64 = lean_ctor_get(x_3, 0); lean_inc(x_64); -x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); -x_66 = lean_ctor_get(x_3, 2); -lean_inc(x_66); +x_65 = lean_ctor_get(x_3, 1); +lean_inc(x_65); +x_66 = lean_ctor_get_uint8(x_3, sizeof(void*)*3); +x_67 = lean_ctor_get(x_3, 2); +lean_inc(x_67); lean_dec_ref(x_3); -x_67 = l_Lean_IR_EmitLLVM_emitDec(x_1, x_2, x_63, x_64, x_65, x_4, x_5); -lean_dec(x_64); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -if (lean_obj_tag(x_68) == 0) +x_68 = l_Lean_IR_EmitLLVM_emitDec(x_1, x_2, x_64, x_65, x_66, x_4, x_5); +lean_dec(x_65); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +if (lean_obj_tag(x_69) == 0) { -lean_dec_ref(x_68); -lean_dec(x_66); +lean_dec_ref(x_69); +lean_dec(x_67); lean_dec_ref(x_5); lean_dec(x_4); -return x_67; +return x_68; } else { +lean_dec_ref(x_69); lean_dec_ref(x_68); -lean_dec_ref(x_67); -x_3 = x_66; +x_3 = x_67; goto _start; } } else { -lean_object* x_70; -x_70 = lean_ctor_get(x_3, 2); -lean_inc(x_70); +lean_object* x_71; +x_71 = lean_ctor_get(x_3, 2); +lean_inc(x_71); lean_dec_ref(x_3); -x_3 = x_70; +x_3 = x_71; goto _start; } } case 8: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_3, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_3, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_3, 0); lean_inc(x_73); +x_74 = lean_ctor_get(x_3, 1); +lean_inc(x_74); lean_dec_ref(x_3); -x_74 = l_Lean_IR_EmitLLVM_emitDel(x_1, x_2, x_72, x_4, x_5); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -if (lean_obj_tag(x_75) == 0) +x_75 = l_Lean_IR_EmitLLVM_emitDel(x_1, x_2, x_73, x_4, x_5); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_obj_tag(x_76) == 0) { -lean_dec_ref(x_75); -lean_dec(x_73); +lean_dec_ref(x_76); +lean_dec(x_74); lean_dec_ref(x_5); lean_dec(x_4); -return x_74; +return x_75; } else { +lean_dec_ref(x_76); lean_dec_ref(x_75); -lean_dec_ref(x_74); -x_3 = x_73; +x_3 = x_74; goto _start; } } case 9: { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_3, 1); -lean_inc(x_77); -x_78 = lean_ctor_get(x_3, 2); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_3, 1); lean_inc(x_78); -x_79 = lean_ctor_get(x_3, 3); -lean_inc_ref(x_79); +x_79 = lean_ctor_get(x_3, 2); +lean_inc(x_79); +x_80 = lean_ctor_get(x_3, 3); +lean_inc_ref(x_80); lean_dec_ref(x_3); -x_80 = l_Lean_IR_EmitLLVM_emitCase(x_1, x_2, x_77, x_78, x_79, x_4, x_5); -return x_80; +x_81 = l_Lean_IR_EmitLLVM_emitCase(x_1, x_2, x_78, x_79, x_80, x_4, x_5); +return x_81; } case 10: { -lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_81 = lean_ctor_get(x_3, 0); -lean_inc(x_81); +lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_82 = lean_ctor_get(x_3, 0); +lean_inc(x_82); lean_dec_ref(x_3); -x_82 = l_Lean_IR_EmitLLVM_emitBlock___closed__0; -x_83 = l_Lean_IR_EmitLLVM_emitArgVal(x_1, x_2, x_81, x_82, x_4, x_5); +x_83 = l_Lean_IR_EmitLLVM_emitBlock___closed__0; +x_84 = l_Lean_IR_EmitLLVM_emitArgVal(x_1, x_2, x_82, x_83, x_4, x_5); lean_dec_ref(x_5); lean_dec(x_4); -x_84 = !lean_is_exclusive(x_83); -if (x_84 == 0) +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_85; -x_85 = lean_ctor_get(x_83, 0); -if (lean_obj_tag(x_85) == 0) +lean_object* x_86; +x_86 = lean_ctor_get(x_84, 0); +if (lean_obj_tag(x_86) == 0) { -uint8_t x_86; -x_86 = !lean_is_exclusive(x_85); -if (x_86 == 0) +uint8_t x_87; +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) { -return x_83; +return x_84; } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_85, 0); -lean_inc(x_87); -lean_dec(x_85); -x_88 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_83, 0, x_88); -return x_83; +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(0, 1, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_84, 0, x_89); +return x_84; } } else { -lean_object* x_89; lean_object* x_90; size_t x_91; size_t x_92; lean_object* x_93; -x_89 = lean_ctor_get(x_85, 0); -lean_inc(x_89); -lean_dec_ref(x_85); -x_90 = lean_ctor_get(x_89, 1); +lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; +x_90 = lean_ctor_get(x_86, 0); lean_inc(x_90); -lean_dec(x_89); -x_91 = lean_unbox_usize(x_90); +lean_dec_ref(x_86); +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); lean_dec(x_90); -x_92 = lean_llvm_build_ret(x_1, x_2, x_91); -x_93 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -lean_ctor_set(x_83, 0, x_93); -return x_83; +x_92 = lean_unbox_usize(x_91); +lean_dec(x_91); +x_93 = lean_llvm_build_ret(x_1, x_2, x_92); +x_94 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +lean_ctor_set(x_84, 0, x_94); +return x_84; } } else { -lean_object* x_94; -x_94 = lean_ctor_get(x_83, 0); -lean_inc(x_94); -lean_dec(x_83); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_ctor_get(x_94, 0); +lean_object* x_95; +x_95 = lean_ctor_get(x_84, 0); lean_inc(x_95); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - x_96 = x_94; +lean_dec(x_84); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + x_97 = x_95; } else { - lean_dec_ref(x_94); - x_96 = lean_box(0); + lean_dec_ref(x_95); + x_97 = lean_box(0); } -if (lean_is_scalar(x_96)) { - x_97 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(0, 1, 0); } else { - x_97 = x_96; + x_98 = x_97; } -lean_ctor_set(x_97, 0, x_95); -x_98 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_98, 0, x_97); -return x_98; +lean_ctor_set(x_98, 0, x_96); +x_99 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_99, 0, x_98); +return x_99; } else { -lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; lean_object* x_103; lean_object* x_104; -x_99 = lean_ctor_get(x_94, 0); -lean_inc(x_99); -lean_dec_ref(x_94); -x_100 = lean_ctor_get(x_99, 1); +lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; lean_object* x_104; lean_object* x_105; +x_100 = lean_ctor_get(x_95, 0); lean_inc(x_100); -lean_dec(x_99); -x_101 = lean_unbox_usize(x_100); +lean_dec_ref(x_95); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); lean_dec(x_100); -x_102 = lean_llvm_build_ret(x_1, x_2, x_101); -x_103 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; -x_104 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_104, 0, x_103); -return x_104; +x_102 = lean_unbox_usize(x_101); +lean_dec(x_101); +x_103 = lean_llvm_build_ret(x_1, x_2, x_102); +x_104 = l_Lean_IR_EmitLLVM_addVarToState___redArg___closed__0; +x_105 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_105, 0, x_104); +return x_105; } } } case 11: { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_3, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_106); +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_3, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_3, 1); +lean_inc_ref(x_107); lean_dec_ref(x_3); -x_107 = l_Lean_IR_EmitLLVM_emitJmp(x_1, x_2, x_105, x_106, x_4, x_5); +x_108 = l_Lean_IR_EmitLLVM_emitJmp(x_1, x_2, x_106, x_107, x_4, x_5); lean_dec_ref(x_5); lean_dec(x_4); -lean_dec_ref(x_106); -return x_107; +lean_dec_ref(x_107); +return x_108; } default: { -lean_object* x_108; +lean_object* x_109; lean_dec(x_4); -x_108 = l_Lean_IR_EmitLLVM_emitUnreachable___redArg(x_1, x_2, x_5); +x_109 = l_Lean_IR_EmitLLVM_emitUnreachable___redArg(x_1, x_2, x_5); lean_dec_ref(x_5); -return x_108; +return x_109; } } } @@ -32865,7 +33679,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; uint8_t x_38; size_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; size_t x_80; lean_object* 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); @@ -32893,7 +33707,7 @@ lean_object* x_43; lean_object* x_44; lean_object* x_45; size_t x_46; lean_objec x_43 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__5; x_44 = lean_st_ref_set(x_40, x_43); x_45 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__6; -x_46 = lean_llvm_append_basic_block_in_context(x_1, x_38, x_45); +x_46 = lean_llvm_append_basic_block_in_context(x_1, x_39, x_45); x_47 = lean_box_usize(x_1); x_48 = lean_llvm_position_builder_at_end(x_47, x_3, x_46); x_49 = !lean_is_exclusive(x_41); @@ -32909,7 +33723,7 @@ lean_ctor_set(x_41, 4, x_25); lean_ctor_set(x_41, 3, x_24); lean_inc_ref(x_41); lean_inc(x_40); -x_52 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_39, x_38, x_25, x_40, x_41); +x_52 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_38, x_39, x_25, x_40, x_41); lean_dec_ref(x_25); if (lean_obj_tag(x_52) == 0) { @@ -32965,7 +33779,7 @@ lean_ctor_set(x_59, 4, x_25); lean_ctor_set_usize(x_59, 5, x_58); lean_inc_ref(x_59); lean_inc(x_40); -x_60 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_39, x_38, x_25, x_40, x_59); +x_60 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_38, x_39, x_25, x_40, x_59); lean_dec_ref(x_25); if (lean_obj_tag(x_60) == 0) { @@ -33004,10 +33818,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); lean_dec(x_64); if (x_74 == 0) @@ -33015,8 +33829,8 @@ if (x_74 == 0) uint64_t x_75; lean_object* x_76; x_75 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__0; x_76 = lean_llvm_set_dll_storage_class(x_1, x_73, x_75); -x_38 = x_73; -x_39 = x_67; +x_38 = x_67; +x_39 = x_73; x_40 = x_69; x_41 = x_70; x_42 = lean_box(0); @@ -33027,8 +33841,8 @@ else uint64_t x_77; lean_object* x_78; x_77 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__1; x_78 = lean_llvm_set_visibility(x_1, x_73, x_77); -x_38 = x_73; -x_39 = x_67; +x_38 = x_67; +x_39 = x_73; x_40 = x_69; x_41 = x_70; x_42 = lean_box(0); @@ -33044,7 +33858,7 @@ x_85 = lean_array_size(x_25); x_86 = 0; lean_inc_ref(x_6); lean_inc(x_5); -x_87 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_85, x_86, x_80, x_5, x_6); +x_87 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_85, x_86, x_83, x_5, x_6); if (lean_obj_tag(x_87) == 0) { uint8_t x_88; @@ -33087,8 +33901,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_80; +x_66 = x_82; x_67 = x_84; x_68 = x_93; x_69 = x_5; @@ -33138,8 +33952,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_80; +x_66 = x_82; x_67 = x_84; x_68 = x_99; x_69 = x_5; @@ -33179,14 +33993,14 @@ return x_102; else { size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec_ref(x_80); +lean_dec_ref(x_83); x_103 = l_LLVM_voidPtrType(x_1); 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_80; +x_66 = x_82; x_67 = x_84; x_68 = x_107; x_69 = x_5; @@ -33254,10 +34068,10 @@ size_t x_120; lean_dec(x_22); x_120 = lean_unbox_usize(x_116); lean_dec(x_116); -x_80 = x_117; +x_80 = x_120; x_81 = lean_box(0); x_82 = x_109; -x_83 = x_120; +x_83 = x_117; x_84 = x_119; goto block_108; } @@ -33268,10 +34082,10 @@ x_121 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_22); lean_dec(x_22); x_122 = lean_unbox_usize(x_116); lean_dec(x_116); -x_80 = x_117; +x_80 = x_122; x_81 = lean_box(0); x_82 = x_109; -x_83 = x_122; +x_83 = x_117; x_84 = x_121; goto block_108; } @@ -33328,10 +34142,10 @@ size_t x_132; lean_dec(x_22); x_132 = lean_unbox_usize(x_128); lean_dec(x_128); -x_80 = x_129; +x_80 = x_132; x_81 = lean_box(0); x_82 = x_109; -x_83 = x_132; +x_83 = x_129; x_84 = x_131; goto block_108; } @@ -33342,10 +34156,10 @@ x_133 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_22); lean_dec(x_22); x_134 = lean_unbox_usize(x_128); lean_dec(x_128); -x_80 = x_129; +x_80 = x_134; x_81 = lean_box(0); x_82 = x_109; -x_83 = x_134; +x_83 = x_129; x_84 = x_133; goto block_108; } @@ -33419,7 +34233,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; uint8_t x_149; size_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; size_t x_186; lean_object* 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); @@ -33446,7 +34260,7 @@ lean_object* x_154; lean_object* x_155; lean_object* x_156; size_t x_157; lean_o x_154 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__5; x_155 = lean_st_ref_set(x_151, x_154); x_156 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__6; -x_157 = lean_llvm_append_basic_block_in_context(x_1, x_149, x_156); +x_157 = lean_llvm_append_basic_block_in_context(x_1, x_150, x_156); x_158 = lean_box_usize(x_1); x_159 = lean_llvm_position_builder_at_end(x_158, x_3, x_157); x_160 = lean_ctor_get(x_152, 0); @@ -33481,7 +34295,7 @@ lean_ctor_set(x_165, 4, x_25); lean_ctor_set_usize(x_165, 5, x_163); lean_inc_ref(x_165); lean_inc(x_151); -x_166 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_150, x_149, x_25, x_151, x_165); +x_166 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_149, x_150, x_25, x_151, x_165); lean_dec_ref(x_25); if (lean_obj_tag(x_166) == 0) { @@ -33519,10 +34333,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); lean_dec(x_170); if (x_180 == 0) @@ -33530,8 +34344,8 @@ if (x_180 == 0) uint64_t x_181; lean_object* x_182; x_181 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__0; x_182 = lean_llvm_set_dll_storage_class(x_1, x_179, x_181); -x_149 = x_179; -x_150 = x_173; +x_149 = x_173; +x_150 = x_179; x_151 = x_175; x_152 = x_176; x_153 = lean_box(0); @@ -33542,8 +34356,8 @@ else uint64_t x_183; lean_object* x_184; x_183 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__1; x_184 = lean_llvm_set_visibility(x_1, x_179, x_183); -x_149 = x_179; -x_150 = x_173; +x_149 = x_173; +x_150 = x_179; x_151 = x_175; x_152 = x_176; x_153 = lean_box(0); @@ -33559,7 +34373,7 @@ x_191 = lean_array_size(x_25); x_192 = 0; lean_inc_ref(x_6); lean_inc(x_5); -x_193 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_191, x_192, x_186, x_5, x_6); +x_193 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_191, x_192, x_189, x_5, x_6); if (lean_obj_tag(x_193) == 0) { lean_object* x_194; lean_object* x_195; @@ -33612,8 +34426,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_186; +x_172 = x_188; x_173 = x_190; x_174 = x_200; x_175 = x_5; @@ -33653,14 +34467,14 @@ return x_203; else { size_t x_204; size_t x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; -lean_dec_ref(x_186); +lean_dec_ref(x_189); x_204 = l_LLVM_voidPtrType(x_1); 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_186; +x_172 = x_188; x_173 = x_190; x_174 = x_208; x_175 = x_5; @@ -33738,10 +34552,10 @@ size_t x_222; lean_dec(x_22); x_222 = lean_unbox_usize(x_218); lean_dec(x_218); -x_186 = x_219; +x_186 = x_222; x_187 = lean_box(0); x_188 = x_210; -x_189 = x_222; +x_189 = x_219; x_190 = x_221; goto block_209; } @@ -33752,10 +34566,10 @@ x_223 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_22); lean_dec(x_22); x_224 = lean_unbox_usize(x_218); lean_dec(x_218); -x_186 = x_219; +x_186 = x_224; x_187 = lean_box(0); x_188 = x_210; -x_189 = x_224; +x_189 = x_219; x_190 = x_223; goto block_209; } @@ -33859,7 +34673,7 @@ return x_244; } else { -lean_object* x_245; lean_object* x_246; size_t x_247; uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_268; lean_object* x_269; size_t x_270; uint8_t x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_284; lean_object* x_285; lean_object* x_286; size_t x_287; uint8_t x_288; lean_object* x_308; uint8_t x_327; +lean_object* x_245; lean_object* x_246; uint8_t x_247; size_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_268; size_t x_269; lean_object* x_270; uint8_t x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; size_t x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; uint8_t x_288; lean_object* x_308; uint8_t x_327; lean_dec(x_240); x_245 = lean_ctor_get(x_239, 0); lean_inc(x_245); @@ -33887,7 +34701,7 @@ lean_object* x_252; lean_object* x_253; lean_object* x_254; size_t x_255; lean_o x_252 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__5; x_253 = lean_st_ref_set(x_249, x_252); x_254 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__6; -x_255 = lean_llvm_append_basic_block_in_context(x_1, x_247, x_254); +x_255 = lean_llvm_append_basic_block_in_context(x_1, x_248, x_254); x_256 = lean_box_usize(x_1); x_257 = lean_llvm_position_builder_at_end(x_256, x_3, x_255); x_258 = lean_ctor_get(x_250, 0); @@ -33922,7 +34736,7 @@ lean_ctor_set(x_263, 4, x_25); lean_ctor_set_usize(x_263, 5, x_261); lean_inc_ref(x_263); lean_inc(x_249); -x_264 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_248, x_247, x_25, x_249, x_263); +x_264 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_247, x_248, x_25, x_249, x_263); lean_dec_ref(x_25); if (lean_obj_tag(x_264) == 0) { @@ -33960,10 +34774,10 @@ goto block_15; block_283: { size_t x_276; size_t x_277; uint8_t x_278; -x_276 = lean_llvm_function_type(x_1, x_270, x_272, x_23); +x_276 = lean_llvm_function_type(x_1, x_269, x_272, x_23); lean_dec_ref(x_272); -x_277 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_269, x_276); -lean_dec_ref(x_269); +x_277 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_270, x_276); +lean_dec_ref(x_270); x_278 = lean_nat_dec_eq(x_268, x_246); lean_dec(x_268); if (x_278 == 0) @@ -33971,8 +34785,8 @@ if (x_278 == 0) uint64_t x_279; lean_object* x_280; x_279 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__0; x_280 = lean_llvm_set_dll_storage_class(x_1, x_277, x_279); -x_247 = x_277; -x_248 = x_271; +x_247 = x_271; +x_248 = x_277; x_249 = x_273; x_250 = x_274; x_251 = lean_box(0); @@ -33983,8 +34797,8 @@ else uint64_t x_281; lean_object* x_282; x_281 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__1; x_282 = lean_llvm_set_visibility(x_1, x_277, x_281); -x_247 = x_277; -x_248 = x_271; +x_247 = x_271; +x_248 = x_277; x_249 = x_273; x_250 = x_274; x_251 = lean_box(0); @@ -34000,7 +34814,7 @@ x_289 = lean_array_size(x_25); x_290 = 0; lean_inc_ref(x_237); lean_inc(x_5); -x_291 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_289, x_290, x_284, x_5, x_237); +x_291 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_25, x_289, x_290, x_287, x_5, x_237); if (lean_obj_tag(x_291) == 0) { lean_object* x_292; lean_object* x_293; @@ -34053,8 +34867,8 @@ lean_dec(x_293); x_298 = lean_ctor_get(x_292, 0); lean_inc(x_298); lean_dec_ref(x_292); -x_269 = x_286; -x_270 = x_287; +x_269 = x_284; +x_270 = x_286; x_271 = x_288; x_272 = x_298; x_273 = x_5; @@ -34094,14 +34908,14 @@ return x_301; else { size_t x_302; size_t x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -lean_dec_ref(x_284); +lean_dec_ref(x_287); x_302 = l_LLVM_voidPtrType(x_1); x_303 = lean_llvm_pointer_type(x_1, x_302); x_304 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_305 = lean_box_usize(x_303); x_306 = lean_array_push(x_304, x_305); -x_269 = x_286; -x_270 = x_287; +x_269 = x_284; +x_270 = x_286; x_271 = x_288; x_272 = x_306; x_273 = x_5; @@ -34179,10 +34993,10 @@ size_t x_320; lean_dec(x_22); x_320 = lean_unbox_usize(x_316); lean_dec(x_316); -x_284 = x_317; +x_284 = x_320; x_285 = lean_box(0); x_286 = x_308; -x_287 = x_320; +x_287 = x_317; x_288 = x_319; goto block_307; } @@ -34193,10 +35007,10 @@ x_321 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_22); lean_dec(x_22); x_322 = lean_unbox_usize(x_316); lean_dec(x_316); -x_284 = x_317; +x_284 = x_322; x_285 = lean_box(0); x_286 = x_308; -x_287 = x_322; +x_287 = x_317; x_288 = x_321; goto block_307; } @@ -34368,7 +35182,7 @@ return x_355; } else { -lean_object* x_356; lean_object* x_357; size_t x_358; uint8_t x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_379; lean_object* x_380; size_t x_381; uint8_t x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_395; lean_object* x_396; lean_object* x_397; size_t x_398; uint8_t x_399; lean_object* x_419; uint8_t x_438; +lean_object* x_356; lean_object* x_357; uint8_t x_358; size_t x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_379; size_t x_380; lean_object* x_381; uint8_t x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; size_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; uint8_t x_399; lean_object* x_419; uint8_t x_438; lean_dec(x_351); x_356 = lean_ctor_get(x_350, 0); lean_inc(x_356); @@ -34396,7 +35210,7 @@ lean_object* x_363; lean_object* x_364; lean_object* x_365; size_t x_366; lean_o x_363 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__5; x_364 = lean_st_ref_set(x_360, x_363); x_365 = l_Lean_IR_EmitLLVM_emitDeclAux___closed__6; -x_366 = lean_llvm_append_basic_block_in_context(x_1, x_358, x_365); +x_366 = lean_llvm_append_basic_block_in_context(x_1, x_359, x_365); x_367 = lean_box_usize(x_1); x_368 = lean_llvm_position_builder_at_end(x_367, x_3, x_366); x_369 = lean_ctor_get(x_361, 0); @@ -34431,7 +35245,7 @@ lean_ctor_set(x_374, 4, x_339); lean_ctor_set_usize(x_374, 5, x_372); lean_inc_ref(x_374); lean_inc(x_360); -x_375 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_359, x_358, x_339, x_360, x_374); +x_375 = l_Lean_IR_EmitLLVM_emitFnArgs(x_1, x_3, x_358, x_359, x_339, x_360, x_374); lean_dec_ref(x_339); if (lean_obj_tag(x_375) == 0) { @@ -34469,10 +35283,10 @@ goto block_15; block_394: { size_t x_387; size_t x_388; uint8_t x_389; -x_387 = lean_llvm_function_type(x_1, x_381, x_383, x_337); +x_387 = lean_llvm_function_type(x_1, x_380, x_383, x_337); lean_dec_ref(x_383); -x_388 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_380, x_387); -lean_dec_ref(x_380); +x_388 = l_Lean_IR_LLVM_getOrAddFunction(x_1, x_2, x_381, x_387); +lean_dec_ref(x_381); x_389 = lean_nat_dec_eq(x_379, x_357); lean_dec(x_379); if (x_389 == 0) @@ -34480,8 +35294,8 @@ if (x_389 == 0) uint64_t x_390; lean_object* x_391; x_390 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__0; x_391 = lean_llvm_set_dll_storage_class(x_1, x_388, x_390); -x_358 = x_388; -x_359 = x_382; +x_358 = x_382; +x_359 = x_388; x_360 = x_384; x_361 = x_385; x_362 = lean_box(0); @@ -34492,8 +35306,8 @@ else uint64_t x_392; lean_object* x_393; x_392 = l_Lean_IR_EmitLLVM_emitFnDeclAux___lam__0___closed__1; x_393 = lean_llvm_set_visibility(x_1, x_388, x_392); -x_358 = x_388; -x_359 = x_382; +x_358 = x_382; +x_359 = x_388; x_360 = x_384; x_361 = x_385; x_362 = lean_box(0); @@ -34509,7 +35323,7 @@ x_400 = lean_array_size(x_339); x_401 = 0; lean_inc_ref(x_348); lean_inc(x_5); -x_402 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_339, x_400, x_401, x_395, x_5, x_348); +x_402 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_EmitLLVM_emitDeclAux_spec__0(x_1, x_339, x_400, x_401, x_398, x_5, x_348); if (lean_obj_tag(x_402) == 0) { lean_object* x_403; lean_object* x_404; @@ -34562,8 +35376,8 @@ lean_dec(x_404); x_409 = lean_ctor_get(x_403, 0); lean_inc(x_409); lean_dec_ref(x_403); -x_380 = x_397; -x_381 = x_398; +x_380 = x_395; +x_381 = x_397; x_382 = x_399; x_383 = x_409; x_384 = x_5; @@ -34603,14 +35417,14 @@ return x_412; else { size_t x_413; size_t x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; -lean_dec_ref(x_395); +lean_dec_ref(x_398); x_413 = l_LLVM_voidPtrType(x_1); x_414 = lean_llvm_pointer_type(x_1, x_413); x_415 = l_Lean_IR_EmitLLVM_callLeanBox___redArg___closed__1; x_416 = lean_box_usize(x_414); x_417 = lean_array_push(x_415, x_416); -x_380 = x_397; -x_381 = x_398; +x_380 = x_395; +x_381 = x_397; x_382 = x_399; x_383 = x_417; x_384 = x_5; @@ -34688,10 +35502,10 @@ size_t x_431; lean_dec(x_336); x_431 = lean_unbox_usize(x_427); lean_dec(x_427); -x_395 = x_428; +x_395 = x_431; x_396 = lean_box(0); x_397 = x_419; -x_398 = x_431; +x_398 = x_428; x_399 = x_430; goto block_418; } @@ -34702,10 +35516,10 @@ x_432 = l_Lean_IR_ExplicitBoxing_isBoxedName(x_336); lean_dec(x_336); x_433 = lean_unbox_usize(x_427); lean_dec(x_427); -x_395 = x_428; +x_395 = x_433; x_396 = lean_box(0); x_397 = x_419; -x_398 = x_433; +x_398 = x_428; x_399 = x_432; goto block_418; } @@ -42977,7 +43791,7 @@ return x_2; 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; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_7; lean_object* x_8; size_t x_9; lean_object* x_10; size_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; x_45 = l_Lean_IR_EmitLLVM_toCName___redArg___closed__0; x_46 = l_Lean_IR_EmitLLVM_toCName___redArg___closed__1; x_47 = l_Lean_IR_EmitLLVM_getDecl___redArg(x_46, x_5); @@ -43038,7 +43852,7 @@ if (lean_is_exclusive(x_48)) { } if (lean_obj_tag(x_55) == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; uint8_t x_71; size_t x_72; size_t x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_133; lean_object* x_311; uint8_t x_312; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; lean_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; size_t x_70; lean_object* x_71; size_t x_72; uint8_t x_73; uint8_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_133; lean_object* x_311; uint8_t x_312; x_57 = lean_ctor_get(x_55, 1); lean_inc_ref(x_57); lean_dec_ref(x_55); @@ -43082,7 +43896,7 @@ lean_dec_ref(x_84); x_86 = lean_ctor_get(x_85, 0); lean_inc(x_86); lean_dec(x_85); -x_87 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_66); +x_87 = l_Lean_IR_EmitLLVM_constInt8___redArg(x_1, x_69); x_88 = lean_ctor_get(x_87, 0); lean_inc(x_88); lean_dec_ref(x_87); @@ -43096,7 +43910,7 @@ lean_dec_ref(x_90); x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); lean_dec(x_91); -x_93 = l_Lean_Name_toString(x_92, x_75); +x_93 = l_Lean_Name_toString(x_92, x_73); x_94 = l_Lean_IR_EmitLLVM_emitMainFn___closed__1; x_95 = lean_string_append(x_93, x_94); x_96 = lean_unbox_usize(x_89); @@ -43112,14 +43926,14 @@ lean_inc(x_100); lean_dec(x_99); x_101 = lean_unbox_usize(x_100); lean_dec(x_100); -x_102 = lean_llvm_build_store(x_1, x_3, x_101, x_73); +x_102 = lean_llvm_build_store(x_1, x_3, x_101, x_72); x_103 = l_LLVM_constTrue(x_1); x_104 = l_Lean_IR_EmitLLVM_callLeanSetPanicMessages___redArg(x_1, x_3, x_103, x_77); lean_dec_ref(x_104); x_105 = l_Lean_IR_EmitLLVM_callLeanIOMarkEndInitialization___redArg(x_1, x_3, x_77); lean_dec_ref(x_105); x_106 = l_Lean_IR_EmitLLVM_emitMainFn___closed__2; -x_107 = lean_llvm_build_load2(x_1, x_3, x_72, x_73, x_106); +x_107 = lean_llvm_build_load2(x_1, x_3, x_75, x_72, x_106); x_108 = l_Lean_IR_EmitLLVM_emitMainFn___closed__3; x_109 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_107, x_108, x_77); x_110 = lean_ctor_get(x_109, 0); @@ -43130,24 +43944,24 @@ lean_inc(x_111); lean_dec(x_110); x_112 = lean_box_usize(x_1); x_113 = lean_box_usize(x_107); -x_114 = lean_box_usize(x_64); -x_115 = lean_box_usize(x_61); -x_116 = lean_box_usize(x_63); -x_117 = lean_box_usize(x_69); -x_118 = lean_box_usize(x_68); -lean_inc(x_66); -lean_inc(x_65); +x_114 = lean_box_usize(x_65); +x_115 = lean_box_usize(x_66); +x_116 = lean_box_usize(x_62); +x_117 = lean_box_usize(x_63); +x_118 = lean_box_usize(x_61); +lean_inc(x_69); +lean_inc(x_67); x_119 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__3___boxed), 17, 13); lean_closure_set(x_119, 0, x_112); lean_closure_set(x_119, 1, x_113); lean_closure_set(x_119, 2, x_60); -lean_closure_set(x_119, 3, x_70); +lean_closure_set(x_119, 3, x_68); lean_closure_set(x_119, 4, x_106); lean_closure_set(x_119, 5, x_114); -lean_closure_set(x_119, 6, x_65); +lean_closure_set(x_119, 6, x_67); lean_closure_set(x_119, 7, x_115); lean_closure_set(x_119, 8, x_116); -lean_closure_set(x_119, 9, x_66); +lean_closure_set(x_119, 9, x_69); lean_closure_set(x_119, 10, x_59); lean_closure_set(x_119, 11, x_117); lean_closure_set(x_119, 12, x_118); @@ -43167,10 +43981,10 @@ if (lean_obj_tag(x_123) == 0) lean_dec_ref(x_123); lean_dec_ref(x_77); lean_dec(x_76); -lean_dec_ref(x_74); -lean_dec_ref(x_67); -lean_dec(x_66); -lean_dec(x_65); +lean_dec_ref(x_71); +lean_dec(x_69); +lean_dec(x_67); +lean_dec_ref(x_64); return x_122; } else @@ -43179,16 +43993,16 @@ lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_dec_ref(x_123); lean_dec_ref(x_122); x_124 = lean_box_usize(x_1); -x_125 = lean_box_usize(x_62); -x_126 = lean_box_usize(x_64); +x_125 = lean_box_usize(x_70); +x_126 = lean_box_usize(x_65); x_127 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__4___boxed), 10, 6); lean_closure_set(x_127, 0, x_124); lean_closure_set(x_127, 1, x_125); lean_closure_set(x_127, 2, x_126); lean_closure_set(x_127, 3, x_106); -lean_closure_set(x_127, 4, x_67); -lean_closure_set(x_127, 5, x_66); -x_128 = l_Lean_Environment_find_x3f(x_74, x_46, x_71); +lean_closure_set(x_127, 4, x_64); +lean_closure_set(x_127, 5, x_69); +x_128 = l_Lean_Environment_find_x3f(x_71, x_46, x_74); if (lean_obj_tag(x_128) == 0) { lean_object* x_129; lean_object* x_130; @@ -43196,16 +44010,16 @@ x_129 = l_Lean_IR_EmitLLVM_emitMainFn___closed__8; x_130 = l_panic___at___00Lean_IR_EmitLLVM_emitMainFn_spec__1(x_129); x_7 = x_106; x_8 = x_127; -x_9 = x_62; -x_10 = x_64; -x_11 = x_65; +x_9 = x_65; +x_10 = x_67; +x_11 = x_70; x_12 = x_106; x_13 = x_76; -x_14 = lean_box(0); -x_15 = x_72; -x_16 = x_73; +x_14 = x_72; +x_15 = lean_box(0); +x_16 = x_77; x_17 = x_108; -x_18 = x_77; +x_18 = x_75; x_19 = x_130; goto block_44; } @@ -43217,16 +44031,16 @@ lean_inc(x_131); lean_dec_ref(x_128); x_7 = x_106; x_8 = x_127; -x_9 = x_62; -x_10 = x_64; -x_11 = x_65; +x_9 = x_65; +x_10 = x_67; +x_11 = x_70; x_12 = x_106; x_13 = x_76; -x_14 = lean_box(0); -x_15 = x_72; -x_16 = x_73; +x_14 = x_72; +x_15 = lean_box(0); +x_16 = x_77; x_17 = x_108; -x_18 = x_77; +x_18 = x_75; x_19 = x_131; goto block_44; } @@ -43236,10 +44050,10 @@ else { lean_dec_ref(x_77); lean_dec(x_76); -lean_dec_ref(x_74); -lean_dec_ref(x_67); -lean_dec(x_66); -lean_dec(x_65); +lean_dec_ref(x_71); +lean_dec(x_69); +lean_dec(x_67); +lean_dec_ref(x_64); return x_122; } } @@ -43446,33 +44260,33 @@ if (x_205 == 0) lean_object* x_208; size_t x_209; size_t x_210; size_t x_211; size_t x_212; size_t x_213; size_t x_214; size_t x_215; x_208 = l_Lean_IR_EmitLLVM_callLeanInitializeRuntimeModule___redArg(x_1, x_3, x_5); lean_dec_ref(x_208); -x_209 = lean_unbox_usize(x_180); -lean_dec(x_180); -x_210 = lean_unbox_usize(x_183); -x_211 = lean_unbox_usize(x_193); -x_212 = lean_unbox_usize(x_170); +x_209 = lean_unbox_usize(x_170); lean_dec(x_170); -x_213 = lean_unbox_usize(x_203); +x_210 = lean_unbox_usize(x_203); lean_dec(x_203); -x_214 = lean_unbox_usize(x_183); -lean_dec(x_183); -x_215 = lean_unbox_usize(x_193); +x_211 = lean_unbox_usize(x_193); +x_212 = lean_unbox_usize(x_180); +lean_dec(x_180); +x_213 = lean_unbox_usize(x_183); +x_214 = lean_unbox_usize(x_193); lean_dec(x_193); +x_215 = lean_unbox_usize(x_183); +lean_dec(x_183); x_61 = x_209; -x_62 = x_210; -x_63 = x_195; -x_64 = x_211; -x_65 = x_206; -x_66 = x_207; -x_67 = x_199; -x_68 = x_212; -x_69 = x_213; -x_70 = x_155; -x_71 = x_160; +x_62 = x_195; +x_63 = x_210; +x_64 = x_199; +x_65 = x_211; +x_66 = x_212; +x_67 = x_206; +x_68 = x_155; +x_69 = x_207; +x_70 = x_213; +x_71 = x_143; x_72 = x_214; -x_73 = x_215; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_215; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43483,33 +44297,33 @@ else lean_object* x_216; size_t x_217; size_t x_218; size_t x_219; size_t x_220; size_t x_221; size_t x_222; size_t x_223; x_216 = l_Lean_IR_EmitLLVM_callLeanInitialize___redArg(x_1, x_3, x_5); lean_dec_ref(x_216); -x_217 = lean_unbox_usize(x_180); -lean_dec(x_180); -x_218 = lean_unbox_usize(x_183); -x_219 = lean_unbox_usize(x_193); -x_220 = lean_unbox_usize(x_170); +x_217 = lean_unbox_usize(x_170); lean_dec(x_170); -x_221 = lean_unbox_usize(x_203); +x_218 = lean_unbox_usize(x_203); lean_dec(x_203); -x_222 = lean_unbox_usize(x_183); -lean_dec(x_183); -x_223 = lean_unbox_usize(x_193); +x_219 = lean_unbox_usize(x_193); +x_220 = lean_unbox_usize(x_180); +lean_dec(x_180); +x_221 = lean_unbox_usize(x_183); +x_222 = lean_unbox_usize(x_193); lean_dec(x_193); +x_223 = lean_unbox_usize(x_183); +lean_dec(x_183); x_61 = x_217; -x_62 = x_218; -x_63 = x_195; -x_64 = x_219; -x_65 = x_206; -x_66 = x_207; -x_67 = x_199; -x_68 = x_220; -x_69 = x_221; -x_70 = x_155; -x_71 = x_160; +x_62 = x_195; +x_63 = x_218; +x_64 = x_199; +x_65 = x_219; +x_66 = x_220; +x_67 = x_206; +x_68 = x_155; +x_69 = x_207; +x_70 = x_221; +x_71 = x_143; x_72 = x_222; -x_73 = x_223; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_223; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43582,33 +44396,33 @@ if (x_241 == 0) lean_object* x_244; size_t x_245; size_t x_246; size_t x_247; size_t x_248; size_t x_249; size_t x_250; size_t x_251; x_244 = l_Lean_IR_EmitLLVM_callLeanInitializeRuntimeModule___redArg(x_1, x_3, x_5); lean_dec_ref(x_244); -x_245 = lean_unbox_usize(x_180); -lean_dec(x_180); -x_246 = lean_unbox_usize(x_183); -x_247 = lean_unbox_usize(x_229); -x_248 = lean_unbox_usize(x_170); +x_245 = lean_unbox_usize(x_170); lean_dec(x_170); -x_249 = lean_unbox_usize(x_239); +x_246 = lean_unbox_usize(x_239); lean_dec(x_239); -x_250 = lean_unbox_usize(x_183); -lean_dec(x_183); -x_251 = lean_unbox_usize(x_229); +x_247 = lean_unbox_usize(x_229); +x_248 = lean_unbox_usize(x_180); +lean_dec(x_180); +x_249 = lean_unbox_usize(x_183); +x_250 = lean_unbox_usize(x_229); lean_dec(x_229); +x_251 = lean_unbox_usize(x_183); +lean_dec(x_183); x_61 = x_245; -x_62 = x_246; -x_63 = x_231; -x_64 = x_247; -x_65 = x_242; -x_66 = x_243; -x_67 = x_235; -x_68 = x_248; -x_69 = x_249; -x_70 = x_155; -x_71 = x_160; +x_62 = x_231; +x_63 = x_246; +x_64 = x_235; +x_65 = x_247; +x_66 = x_248; +x_67 = x_242; +x_68 = x_155; +x_69 = x_243; +x_70 = x_249; +x_71 = x_143; x_72 = x_250; -x_73 = x_251; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_251; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43619,33 +44433,33 @@ else lean_object* x_252; size_t x_253; size_t x_254; size_t x_255; size_t x_256; size_t x_257; size_t x_258; size_t x_259; x_252 = l_Lean_IR_EmitLLVM_callLeanInitialize___redArg(x_1, x_3, x_5); lean_dec_ref(x_252); -x_253 = lean_unbox_usize(x_180); -lean_dec(x_180); -x_254 = lean_unbox_usize(x_183); -x_255 = lean_unbox_usize(x_229); -x_256 = lean_unbox_usize(x_170); +x_253 = lean_unbox_usize(x_170); lean_dec(x_170); -x_257 = lean_unbox_usize(x_239); +x_254 = lean_unbox_usize(x_239); lean_dec(x_239); -x_258 = lean_unbox_usize(x_183); -lean_dec(x_183); -x_259 = lean_unbox_usize(x_229); +x_255 = lean_unbox_usize(x_229); +x_256 = lean_unbox_usize(x_180); +lean_dec(x_180); +x_257 = lean_unbox_usize(x_183); +x_258 = lean_unbox_usize(x_229); lean_dec(x_229); +x_259 = lean_unbox_usize(x_183); +lean_dec(x_183); x_61 = x_253; -x_62 = x_254; -x_63 = x_231; -x_64 = x_255; -x_65 = x_242; -x_66 = x_243; -x_67 = x_235; -x_68 = x_256; -x_69 = x_257; -x_70 = x_155; -x_71 = x_160; +x_62 = x_231; +x_63 = x_254; +x_64 = x_235; +x_65 = x_255; +x_66 = x_256; +x_67 = x_242; +x_68 = x_155; +x_69 = x_243; +x_70 = x_257; +x_71 = x_143; x_72 = x_258; -x_73 = x_259; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_259; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43779,33 +44593,33 @@ if (x_291 == 0) lean_object* x_294; size_t x_295; size_t x_296; size_t x_297; size_t x_298; size_t x_299; size_t x_300; size_t x_301; x_294 = l_Lean_IR_EmitLLVM_callLeanInitializeRuntimeModule___redArg(x_1, x_3, x_5); lean_dec_ref(x_294); -x_295 = lean_unbox_usize(x_265); -lean_dec(x_265); -x_296 = lean_unbox_usize(x_268); -x_297 = lean_unbox_usize(x_279); -x_298 = lean_unbox_usize(x_170); +x_295 = lean_unbox_usize(x_170); lean_dec(x_170); -x_299 = lean_unbox_usize(x_289); +x_296 = lean_unbox_usize(x_289); lean_dec(x_289); -x_300 = lean_unbox_usize(x_268); -lean_dec(x_268); -x_301 = lean_unbox_usize(x_279); +x_297 = lean_unbox_usize(x_279); +x_298 = lean_unbox_usize(x_265); +lean_dec(x_265); +x_299 = lean_unbox_usize(x_268); +x_300 = lean_unbox_usize(x_279); lean_dec(x_279); +x_301 = lean_unbox_usize(x_268); +lean_dec(x_268); x_61 = x_295; -x_62 = x_296; -x_63 = x_281; -x_64 = x_297; -x_65 = x_292; -x_66 = x_293; -x_67 = x_285; -x_68 = x_298; -x_69 = x_299; -x_70 = x_155; -x_71 = x_160; +x_62 = x_281; +x_63 = x_296; +x_64 = x_285; +x_65 = x_297; +x_66 = x_298; +x_67 = x_292; +x_68 = x_155; +x_69 = x_293; +x_70 = x_299; +x_71 = x_143; x_72 = x_300; -x_73 = x_301; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_301; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43816,33 +44630,33 @@ else lean_object* x_302; size_t x_303; size_t x_304; size_t x_305; size_t x_306; size_t x_307; size_t x_308; size_t x_309; x_302 = l_Lean_IR_EmitLLVM_callLeanInitialize___redArg(x_1, x_3, x_5); lean_dec_ref(x_302); -x_303 = lean_unbox_usize(x_265); -lean_dec(x_265); -x_304 = lean_unbox_usize(x_268); -x_305 = lean_unbox_usize(x_279); -x_306 = lean_unbox_usize(x_170); +x_303 = lean_unbox_usize(x_170); lean_dec(x_170); -x_307 = lean_unbox_usize(x_289); +x_304 = lean_unbox_usize(x_289); lean_dec(x_289); -x_308 = lean_unbox_usize(x_268); -lean_dec(x_268); -x_309 = lean_unbox_usize(x_279); +x_305 = lean_unbox_usize(x_279); +x_306 = lean_unbox_usize(x_265); +lean_dec(x_265); +x_307 = lean_unbox_usize(x_268); +x_308 = lean_unbox_usize(x_279); lean_dec(x_279); +x_309 = lean_unbox_usize(x_268); +lean_dec(x_268); x_61 = x_303; -x_62 = x_304; -x_63 = x_281; -x_64 = x_305; -x_65 = x_292; -x_66 = x_293; -x_67 = x_285; -x_68 = x_306; -x_69 = x_307; -x_70 = x_155; -x_71 = x_160; +x_62 = x_281; +x_63 = x_304; +x_64 = x_285; +x_65 = x_305; +x_66 = x_306; +x_67 = x_292; +x_68 = x_155; +x_69 = x_293; +x_70 = x_307; +x_71 = x_143; x_72 = x_308; -x_73 = x_309; -x_74 = x_143; -x_75 = x_133; +x_73 = x_133; +x_74 = x_160; +x_75 = x_309; x_76 = x_4; x_77 = x_5; x_78 = lean_box(0); @@ -43874,10 +44688,10 @@ return x_316; block_44: { lean_object* x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; -x_20 = l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg(x_1, x_3, x_18); +x_20 = l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager___redArg(x_1, x_3, x_16); lean_dec_ref(x_20); -x_21 = lean_llvm_build_load2(x_1, x_3, x_15, x_16, x_12); -x_22 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_21, x_17, x_18); +x_21 = lean_llvm_build_load2(x_1, x_3, x_18, x_14, x_12); +x_22 = l_Lean_IR_EmitLLVM_callLeanIOResultIsOk___redArg(x_1, x_3, x_21, x_17, x_16); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec_ref(x_22); @@ -43892,20 +44706,20 @@ x_27 = l_Lean_Expr_appArg_x21(x_26); lean_dec_ref(x_26); x_28 = lean_box_usize(x_1); x_29 = lean_box_usize(x_21); -x_30 = lean_box_usize(x_9); -x_31 = lean_box_usize(x_10); +x_30 = lean_box_usize(x_11); +x_31 = lean_box_usize(x_9); x_32 = lean_alloc_closure((void*)(l_Lean_IR_EmitLLVM_emitMainFn___lam__5___boxed), 11, 7); lean_closure_set(x_32, 0, x_27); lean_closure_set(x_32, 1, x_28); lean_closure_set(x_32, 2, x_29); -lean_closure_set(x_32, 3, x_11); +lean_closure_set(x_32, 3, x_10); lean_closure_set(x_32, 4, x_30); lean_closure_set(x_32, 5, x_31); lean_closure_set(x_32, 6, x_7); x_33 = l_Lean_IR_EmitLLVM_emitMainFn___closed__0; x_34 = lean_unbox_usize(x_24); lean_dec(x_24); -x_35 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_33, x_34, x_32, x_8, x_13, x_18); +x_35 = l_Lean_IR_EmitLLVM_buildIfThenElse__(x_1, x_3, x_33, x_34, x_32, x_8, x_13, x_16); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; @@ -44781,7 +45595,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean x_13 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_emitLLVM_spec__3___closed__0; x_14 = lean_string_append(x_13, x_11); lean_dec_ref(x_11); -x_15 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; +x_15 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -44844,7 +45658,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean x_13 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_IR_emitLLVM_spec__4___closed__0; x_14 = lean_string_append(x_13, x_11); lean_dec_ref(x_11); -x_15 = l_Lean_IR_EmitLLVM_emitSProj___closed__8; +x_15 = l_Lean_IR_EmitLLVM_emitSProj___closed__12; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -45623,8 +46437,10 @@ l_Lean_IR_EmitLLVM_emitApp___closed__2 = _init_l_Lean_IR_EmitLLVM_emitApp___clos lean_mark_persistent(l_Lean_IR_EmitLLVM_emitApp___closed__2); l_Lean_IR_EmitLLVM_callLeanCtorGet___redArg___closed__0 = _init_l_Lean_IR_EmitLLVM_callLeanCtorGet___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_callLeanCtorGet___redArg___closed__0); -l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0 = _init_l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0(); -lean_mark_persistent(l_Lean_IR_EmitLLVM_callLeanCtorGetUsize___redArg___closed__0); +l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0 = _init_l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__0); +l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1 = _init_l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_callLeanCtorGetPtrSizedType___redArg___closed__1); l_Lean_IR_EmitLLVM_emitSProj___closed__0 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__0); l_Lean_IR_EmitLLVM_emitSProj___closed__1 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__1(); @@ -45643,6 +46459,14 @@ l_Lean_IR_EmitLLVM_emitSProj___closed__7 = _init_l_Lean_IR_EmitLLVM_emitSProj___ lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__7); l_Lean_IR_EmitLLVM_emitSProj___closed__8 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__8(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__8); +l_Lean_IR_EmitLLVM_emitSProj___closed__9 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__9); +l_Lean_IR_EmitLLVM_emitSProj___closed__10 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__10(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__10); +l_Lean_IR_EmitLLVM_emitSProj___closed__11 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__11(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__11); +l_Lean_IR_EmitLLVM_emitSProj___closed__12 = _init_l_Lean_IR_EmitLLVM_emitSProj___closed__12(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSProj___closed__12); l_Lean_IR_EmitLLVM_callLeanIsExclusive___redArg___closed__0 = _init_l_Lean_IR_EmitLLVM_callLeanIsExclusive___redArg___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_callLeanIsExclusive___redArg___closed__0); l_Lean_IR_EmitLLVM_callLeanIsScalar___redArg___closed__0 = _init_l_Lean_IR_EmitLLVM_callLeanIsScalar___redArg___closed__0(); @@ -45657,6 +46481,12 @@ l_Lean_IR_EmitLLVM_emitBox___closed__3 = _init_l_Lean_IR_EmitLLVM_emitBox___clos lean_mark_persistent(l_Lean_IR_EmitLLVM_emitBox___closed__3); l_Lean_IR_EmitLLVM_emitBox___closed__4 = _init_l_Lean_IR_EmitLLVM_emitBox___closed__4(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitBox___closed__4); +l_Lean_IR_EmitLLVM_emitBox___closed__5 = _init_l_Lean_IR_EmitLLVM_emitBox___closed__5(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitBox___closed__5); +l_Lean_IR_EmitLLVM_emitBox___closed__6 = _init_l_Lean_IR_EmitLLVM_emitBox___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitBox___closed__6); +l_Lean_IR_EmitLLVM_emitBox___closed__7 = _init_l_Lean_IR_EmitLLVM_emitBox___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitBox___closed__7); l_Lean_IR_EmitLLVM_callUnboxForType___closed__0 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__0); l_Lean_IR_EmitLLVM_callUnboxForType___closed__1 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__1(); @@ -45669,6 +46499,12 @@ l_Lean_IR_EmitLLVM_callUnboxForType___closed__4 = _init_l_Lean_IR_EmitLLVM_callU lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__4); l_Lean_IR_EmitLLVM_callUnboxForType___closed__5 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__5(); lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__5); +l_Lean_IR_EmitLLVM_callUnboxForType___closed__6 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__6(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__6); +l_Lean_IR_EmitLLVM_callUnboxForType___closed__7 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__7); +l_Lean_IR_EmitLLVM_callUnboxForType___closed__8 = _init_l_Lean_IR_EmitLLVM_callUnboxForType___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_callUnboxForType___closed__8); l_Lean_IR_EmitLLVM_emitReset___lam__0___closed__0 = _init_l_Lean_IR_EmitLLVM_emitReset___lam__0___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitReset___lam__0___closed__0); l_Lean_IR_EmitLLVM_emitReset___lam__1___closed__0 = _init_l_Lean_IR_EmitLLVM_emitReset___lam__1___closed__0(); @@ -45685,6 +46521,8 @@ l_Lean_IR_EmitLLVM_emitTag___closed__1 = _init_l_Lean_IR_EmitLLVM_emitTag___clos lean_mark_persistent(l_Lean_IR_EmitLLVM_emitTag___closed__1); l_Lean_IR_EmitLLVM_emitUSet___closed__0 = _init_l_Lean_IR_EmitLLVM_emitUSet___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitUSet___closed__0); +l_Lean_IR_EmitLLVM_emitUSet___closed__1 = _init_l_Lean_IR_EmitLLVM_emitUSet___closed__1(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitUSet___closed__1); l_Lean_IR_EmitLLVM_emitTailCall___closed__0 = _init_l_Lean_IR_EmitLLVM_emitTailCall___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitTailCall___closed__0); l_Lean_IR_EmitLLVM_emitTailCall___closed__1 = _init_l_Lean_IR_EmitLLVM_emitTailCall___closed__1(); @@ -45711,6 +46549,14 @@ l_Lean_IR_EmitLLVM_emitSSet___closed__5 = _init_l_Lean_IR_EmitLLVM_emitSSet___cl lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__5); l_Lean_IR_EmitLLVM_emitSSet___closed__6 = _init_l_Lean_IR_EmitLLVM_emitSSet___closed__6(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__6); +l_Lean_IR_EmitLLVM_emitSSet___closed__7 = _init_l_Lean_IR_EmitLLVM_emitSSet___closed__7(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__7); +l_Lean_IR_EmitLLVM_emitSSet___closed__8 = _init_l_Lean_IR_EmitLLVM_emitSSet___closed__8(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__8); +l_Lean_IR_EmitLLVM_emitSSet___closed__9 = _init_l_Lean_IR_EmitLLVM_emitSSet___closed__9(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__9); +l_Lean_IR_EmitLLVM_emitSSet___closed__10 = _init_l_Lean_IR_EmitLLVM_emitSSet___closed__10(); +lean_mark_persistent(l_Lean_IR_EmitLLVM_emitSSet___closed__10); l_Lean_IR_EmitLLVM_emitDel___closed__0 = _init_l_Lean_IR_EmitLLVM_emitDel___closed__0(); lean_mark_persistent(l_Lean_IR_EmitLLVM_emitDel___closed__0); l_Lean_IR_EmitLLVM_ensureHasDefault_x27___closed__0 = _init_l_Lean_IR_EmitLLVM_ensureHasDefault_x27___closed__0(); diff --git a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c index ec5b02ea0515..604990ffde11 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ExpandResetReuse.c @@ -2710,174 +2710,168 @@ goto _start; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; uint8_t x_38; 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_inc(x_36); +x_36 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); +x_37 = lean_ctor_get(x_2, 3); +lean_inc(x_37); 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) -{ -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; +x_38 = l_Lean_IR_ExpandResetReuse_isSelfUSet(x_1, x_33, x_34, x_35); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_37); +x_40 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 3, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_36); +return x_40; } else { lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); -x_2 = x_36; +x_2 = x_37; 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) -{ -lean_object* x_49; -x_49 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_47); -lean_ctor_set(x_2, 5, x_49); +uint8_t x_42; +x_42 = !lean_is_exclusive(x_2); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_43 = lean_ctor_get(x_2, 0); +x_44 = lean_ctor_get(x_2, 1); +x_45 = lean_ctor_get(x_2, 2); +x_46 = lean_ctor_get(x_2, 3); +x_47 = lean_ctor_get(x_2, 4); +x_48 = lean_ctor_get(x_2, 5); +x_49 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_43, x_44, x_45, x_46); +if (x_49 == 0) +{ +lean_object* x_50; +x_50 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_48); +lean_ctor_set(x_2, 5, x_50); return x_2; } else { lean_free_object(x_2); +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_48; 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_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_52 = lean_ctor_get(x_2, 0); +x_53 = lean_ctor_get(x_2, 1); +x_54 = lean_ctor_get(x_2, 2); +x_55 = lean_ctor_get(x_2, 3); +x_56 = lean_ctor_get(x_2, 4); +x_57 = lean_ctor_get(x_2, 5); +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_58 = l_Lean_IR_ExpandResetReuse_isSelfSSet(x_1, x_52, x_53, x_54, x_55); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +x_59 = l_Lean_IR_ExpandResetReuse_removeSelfSet(x_1, x_57); +x_60 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_60, 0, x_52); +lean_ctor_set(x_60, 1, x_53); +lean_ctor_set(x_60, 2, x_54); +lean_ctor_set(x_60, 3, x_55); +lean_ctor_set(x_60, 4, x_56); +lean_ctor_set(x_60, 5, x_59); +return x_60; } else { +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_57; goto _start; } } } case 9: { -uint8_t x_61; -x_61 = !lean_is_exclusive(x_2); -if (x_61 == 0) +uint8_t x_62; +x_62 = !lean_is_exclusive(x_2); +if (x_62 == 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_63; size_t x_64; size_t x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_2, 3); +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_removeSelfSet_spec__0(x_1, x_64, x_65, x_63); +lean_ctor_set(x_2, 3, x_66); 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_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_2, 0); +x_68 = lean_ctor_get(x_2, 1); +x_69 = lean_ctor_get(x_2, 2); +x_70 = lean_ctor_get(x_2, 3); +lean_inc(x_70); lean_inc(x_69); lean_inc(x_68); lean_inc(x_67); -lean_inc(x_66); 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_71 = lean_array_size(x_70); +x_72 = 0; +x_73 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExpandResetReuse_removeSelfSet_spec__0(x_1, x_71, x_72, x_70); +x_74 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_74, 0, x_67); +lean_ctor_set(x_74, 1, x_68); +lean_ctor_set(x_74, 2, x_69); +lean_ctor_set(x_74, 3, x_73); +return x_74; } } default: { -uint8_t x_74; -x_74 = l_Lean_IR_FnBody_isTerminal(x_2); -if (x_74 == 0) +uint8_t x_75; +x_75 = l_Lean_IR_FnBody_isTerminal(x_2); +if (x_75 == 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; -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; goto block_8; } -case 2: +case 1: { lean_object* x_77; x_77 = lean_ctor_get(x_2, 3); @@ -2885,39 +2879,39 @@ lean_inc(x_77); x_3 = x_77; goto block_8; } -case 3: +case 2: { lean_object* x_78; -x_78 = lean_ctor_get(x_2, 2); +x_78 = lean_ctor_get(x_2, 3); lean_inc(x_78); x_3 = x_78; goto block_8; } -case 4: +case 3: { lean_object* x_79; -x_79 = lean_ctor_get(x_2, 3); +x_79 = lean_ctor_get(x_2, 2); lean_inc(x_79); x_3 = x_79; goto block_8; } -case 5: +case 4: { lean_object* x_80; -x_80 = lean_ctor_get(x_2, 5); +x_80 = lean_ctor_get(x_2, 3); lean_inc(x_80); x_3 = x_80; goto block_8; } -case 6: +case 5: { lean_object* x_81; -x_81 = lean_ctor_get(x_2, 2); +x_81 = lean_ctor_get(x_2, 5); lean_inc(x_81); x_3 = x_81; goto block_8; } -case 7: +case 6: { lean_object* x_82; x_82 = lean_ctor_get(x_2, 2); @@ -2925,14 +2919,22 @@ lean_inc(x_82); x_3 = x_82; goto block_8; } -case 8: +case 7: { lean_object* x_83; -x_83 = lean_ctor_get(x_2, 1); +x_83 = lean_ctor_get(x_2, 2); lean_inc(x_83); x_3 = x_83; goto block_8; } +case 8: +{ +lean_object* x_84; +x_84 = lean_ctor_get(x_2, 1); +lean_inc(x_84); +x_3 = x_84; +goto block_8; +} default: { lean_inc(x_2); diff --git a/stage0/stdlib/Lean/Compiler/IR/Format.c b/stage0/stdlib/Lean/Compiler/IR/Format.c index 94562262c766..7126bbfdc2ae 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Format.c +++ b/stage0/stdlib/Lean/Compiler/IR/Format.c @@ -95,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*); @@ -110,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*); @@ -118,6 +120,7 @@ static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInf LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__11; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__18; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__0; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; LEAN_EXPORT lean_object* l_Lean_IR_instToFormatParam___private__1(lean_object*); @@ -125,6 +128,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_formatArray___at___00Lean_IR_formatParams_spe static lean_object* l_Lean_IR_formatArray___redArg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(lean_object*); static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__20; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39; static lean_object* l_Lean_IR_formatAlt___closed__3; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__20; @@ -139,6 +143,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; @@ -158,6 +163,7 @@ static lean_object* l_Lean_IR_formatFnBodyHead___closed__7; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__17; static lean_object* l_Lean_IR_formatFnBodyHead___closed__21; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__1; +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__21; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__8; static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__3; @@ -188,17 +194,23 @@ 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___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40; 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_formatIRType___closed__38; 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; static lean_object* l_Lean_IR_instToFormatArg___closed__0; LEAN_EXPORT lean_object* l_Lean_IR_instToStringIRType; 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*); +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37; static lean_object* l_Lean_IR_formatFnBodyHead___closed__26; LEAN_EXPORT lean_object* l_Lean_IR_formatFnBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType_spec__0(lean_object*, lean_object*); +static lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29; LEAN_EXPORT lean_object* l_Lean_IR_instToFormatIRType___private__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_instToStringExpr___lam__0(lean_object*); static lean_object* l_Lean_IR_formatFnBodyHead___closed__8; @@ -984,25 +996,23 @@ 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(", u", 3, 3); return x_1; } } static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__14; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", s", 3, 3); +return x_1; } } static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__16() { _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; } } @@ -1020,7 +1030,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("pap ", 4, 4); +x_1 = lean_mk_string_unchecked(", ", 2, 2); return x_1; } } @@ -1038,7 +1048,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("app ", 4, 4); +x_1 = lean_mk_string_unchecked("pap ", 4, 4); return x_1; } } @@ -1056,7 +1066,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("box ", 4, 4); +x_1 = lean_mk_string_unchecked("app ", 4, 4); return x_1; } } @@ -1074,7 +1084,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("unbox ", 6, 6); +x_1 = lean_mk_string_unchecked("box ", 4, 4); return x_1; } } @@ -1092,7 +1102,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("isShared ", 9, 9); +x_1 = lean_mk_string_unchecked("unbox ", 6, 6); return x_1; } } @@ -1106,6 +1116,24 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isShared ", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr(lean_object* x_1) { _start: { @@ -1333,374 +1361,366 @@ 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); +lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_88 = lean_ctor_get(x_1, 0); +lean_inc(x_88); +x_89 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_90 = lean_ctor_get(x_1, 1); +lean_inc(x_90); +lean_dec_ref(x_1); x_91 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__13; -x_92 = l_Nat_reprFast(x_89); +x_92 = l_Nat_reprFast(x_88); 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; +x_94 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +if (x_89 == 0) +{ +lean_object* x_106; +x_106 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__14; +x_95 = x_106; +goto block_105; } 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; +lean_object* x_107; +x_107 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; +x_95 = x_107; +goto block_105; +} +block_105: +{ +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_96 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_96, 0, x_95); +x_97 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_96); +x_98 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; +x_99 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_101 = l_Nat_reprFast(x_90); +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); +return x_104; } } 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_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; +x_108 = lean_ctor_get(x_1, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_1, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_1, 2); +lean_inc(x_110); 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_111 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__17; +x_112 = l_Nat_reprFast(x_108); +x_113 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_113, 0, x_112); +x_114 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_113); +x_115 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; +x_116 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +x_117 = l_Nat_reprFast(x_109); +x_118 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_118, 0, x_117); +x_119 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_118); +x_120 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__3; +x_121 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_123 = l_Nat_reprFast(x_110); +x_124 = lean_string_append(x_122, x_123); +lean_dec_ref(x_123); +x_125 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_125, 0, x_124); +x_126 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_126, 0, x_121); +lean_ctor_set(x_126, 1, x_125); +return x_126; } case 6: { -uint8_t x_133; -x_133 = !lean_is_exclusive(x_1); -if (x_133 == 0) +uint8_t x_127; +x_127 = !lean_is_exclusive(x_1); +if (x_127 == 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; +lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_128 = lean_ctor_get(x_1, 0); +x_129 = lean_ctor_get(x_1, 1); +x_130 = 1; +x_131 = l_Lean_Name_toString(x_128, x_130); +x_132 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_132, 0, x_131); +x_133 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_129); +lean_dec_ref(x_129); +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_133); +lean_ctor_set(x_1, 0, x_132); +return x_1; +} +else +{ +lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; x_134 = lean_ctor_get(x_1, 0); x_135 = lean_ctor_get(x_1, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_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); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_139); -lean_ctor_set(x_1, 0, x_138); -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_dec(x_1); -x_142 = 1; -x_143 = l_Lean_Name_toString(x_140, 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); -x_146 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; +x_140 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } case 7: { -uint8_t x_147; -x_147 = !lean_is_exclusive(x_1); -if (x_147 == 0) -{ -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); +uint8_t x_141; +x_141 = !lean_is_exclusive(x_1); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_142 = lean_ctor_get(x_1, 0); +x_143 = lean_ctor_get(x_1, 1); +x_144 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__21; +x_145 = 1; +x_146 = l_Lean_Name_toString(x_142, x_145); +x_147 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_147, 0, x_146); 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); -x_155 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_155, 0, x_1); -lean_ctor_set(x_155, 1, x_154); -return x_155; +lean_ctor_set(x_1, 1, x_147); +lean_ctor_set(x_1, 0, x_144); +x_148 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_143); +lean_dec_ref(x_143); +x_149 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_149, 0, x_1); +lean_ctor_set(x_149, 1, x_148); +return x_149; } 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_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_150 = lean_ctor_get(x_1, 0); +x_151 = lean_ctor_get(x_1, 1); +lean_inc(x_151); +lean_inc(x_150); 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; +x_152 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__21; +x_153 = 1; +x_154 = l_Lean_Name_toString(x_150, x_153); +x_155 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_155, 0, x_154); +x_156 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_156, 0, x_152); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_151); +lean_dec_ref(x_151); +x_158 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; } } case 8: { -uint8_t x_165; -x_165 = !lean_is_exclusive(x_1); -if (x_165 == 0) -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_166 = lean_ctor_get(x_1, 0); -x_167 = lean_ctor_get(x_1, 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); -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); +uint8_t x_159; +x_159 = !lean_is_exclusive(x_1); +if (x_159 == 0) +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_160 = lean_ctor_get(x_1, 0); +x_161 = lean_ctor_get(x_1, 1); +x_162 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; +x_163 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_164 = l_Nat_reprFast(x_160); +x_165 = lean_string_append(x_163, x_164); +lean_dec_ref(x_164); +x_166 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_166, 0, x_165); 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); -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; +lean_ctor_set(x_1, 1, x_166); +lean_ctor_set(x_1, 0, x_162); +x_167 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_161); +lean_dec_ref(x_161); +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_1); +lean_ctor_set(x_168, 1, x_167); +return x_168; } 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_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_169 = lean_ctor_get(x_1, 0); +x_170 = lean_ctor_get(x_1, 1); +lean_inc(x_170); +lean_inc(x_169); 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_171 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__23; +x_172 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_173 = l_Nat_reprFast(x_169); +x_174 = lean_string_append(x_172, x_173); +lean_dec_ref(x_173); +x_175 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_175, 0, x_174); +x_176 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_176, 0, x_171); +lean_ctor_set(x_176, 1, x_175); +x_177 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_170); +lean_dec_ref(x_170); +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; } } 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_179; +x_179 = !lean_is_exclusive(x_1); +if (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_1, 1); +x_181 = lean_ctor_get(x_1, 0); +lean_dec(x_181); +x_182 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; +x_183 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_184 = l_Nat_reprFast(x_180); +x_185 = lean_string_append(x_183, x_184); +lean_dec_ref(x_184); +x_186 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_186, 0, x_185); 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_186); +lean_ctor_set(x_1, 0, x_182); 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_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_187 = lean_ctor_get(x_1, 1); +lean_inc(x_187); 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_188 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; +x_189 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_190 = l_Nat_reprFast(x_187); +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); +x_193 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_193, 0, x_188); +lean_ctor_set(x_193, 1, x_192); +return x_193; } } case 10: { -uint8_t x_200; -x_200 = !lean_is_exclusive(x_1); -if (x_200 == 0) +uint8_t x_194; +x_194 = !lean_is_exclusive(x_1); +if (x_194 == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_195 = lean_ctor_get(x_1, 0); +x_196 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; +x_197 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_198 = l_Nat_reprFast(x_195); +x_199 = lean_string_append(x_197, x_198); +lean_dec_ref(x_198); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_199); +x_200 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_200, 0, x_196); +lean_ctor_set(x_200, 1, x_1); +return x_200; +} +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_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; x_201 = lean_ctor_get(x_1, 0); -x_202 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__25; +lean_inc(x_201); +lean_dec(x_1); +x_202 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; 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_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; -} -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_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_206 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_206, 0, x_205); +x_207 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_207, 0, x_202); +lean_ctor_set(x_207, 1, x_206); +return x_207; } } 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_208; lean_object* x_209; +x_208 = lean_ctor_get(x_1, 0); +lean_inc_ref(x_208); lean_dec_ref(x_1); -x_215 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatLitVal(x_214); -return x_215; +x_209 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatLitVal(x_208); +return x_209; } default: { -uint8_t x_216; -x_216 = !lean_is_exclusive(x_1); -if (x_216 == 0) +uint8_t x_210; +x_210 = !lean_is_exclusive(x_1); +if (x_210 == 0) +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_211 = lean_ctor_get(x_1, 0); +x_212 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29; +x_213 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_214 = l_Nat_reprFast(x_211); +x_215 = lean_string_append(x_213, x_214); +lean_dec_ref(x_214); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_215); +x_216 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_216, 0, x_212); +lean_ctor_set(x_216, 1, x_1); +return x_216; +} +else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; x_217 = lean_ctor_get(x_1, 0); -x_218 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27; +lean_inc(x_217); +lean_dec(x_1); +x_218 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29; 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); -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; -} -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_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_222 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_222, 0, x_221); +x_223 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_223, 0, x_218); +lean_ctor_set(x_223, 1, x_222); +return x_223; } } } @@ -2165,6 +2185,96 @@ 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__32() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("i8", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__32; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("i16", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__34; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("i32", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__36; +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__38() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("i64", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__38; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isize", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(lean_object* x_1) { _start: { @@ -2241,7 +2351,7 @@ 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_17 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; 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; @@ -2272,7 +2382,7 @@ 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_30 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; 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; @@ -2308,7 +2418,7 @@ 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_46 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; 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; @@ -2339,7 +2449,7 @@ lean_inc(x_56); 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_59 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__19; 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; @@ -2369,12 +2479,42 @@ lean_object* x_70; x_70 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__29; return x_70; } -default: +case 13: { lean_object* x_71; x_71 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__31; return x_71; } +case 14: +{ +lean_object* x_72; +x_72 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__33; +return x_72; +} +case 15: +{ +lean_object* x_73; +x_73 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__35; +return x_73; +} +case 16: +{ +lean_object* x_74; +x_74 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37; +return x_74; +} +case 17: +{ +lean_object* x_75; +x_75 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39; +return x_75; +} +default: +{ +lean_object* x_76; +x_76 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41; +return x_76; +} } } } @@ -3284,294 +3424,299 @@ 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; 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; 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_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_uint8(x_1, sizeof(void*)*4); 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_78 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_78, 0, x_77); +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); +if (x_63 == 0) +{ +lean_object* x_85; +x_85 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__14; +x_75 = x_85; +goto block_84; +} +else +{ +lean_object* x_86; +x_86 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; +x_75 = x_86; +goto block_84; +} +block_84: +{ +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_76 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Lean_IR_formatFnBodyHead___closed__10; x_79 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 0, x_77); lean_ctor_set(x_79, 1, x_78); -return x_79; +x_80 = l_Nat_reprFast(x_62); +x_81 = lean_string_append(x_65, x_80); +lean_dec_ref(x_80); +x_82 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_82, 0, x_81); +x_83 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_83, 0, x_79); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} } 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_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; +x_87 = lean_ctor_get(x_1, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_1, 1); +lean_inc(x_88); +x_89 = lean_ctor_get(x_1, 2); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 3); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 4); +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_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_92 = l_Lean_IR_formatFnBodyHead___closed__16; +x_93 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_94 = l_Nat_reprFast(x_87); +x_95 = lean_string_append(x_93, x_94); +lean_dec_ref(x_94); +x_96 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_96, 0, x_95); x_97 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 0, x_92); 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_formatCtorInfo___closed__1; +x_99 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Nat_reprFast(x_88); +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_99); 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__19; 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; +x_105 = l_Nat_reprFast(x_89); +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_Lean_IR_formatFnBodyHead___closed__18; +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___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_91); +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_Lean_IR_formatFnBodyHead___closed__3; +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_Nat_reprFast(x_90); +x_115 = lean_string_append(x_93, x_114); +lean_dec_ref(x_114); +x_116 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_116, 0, x_115); +x_117 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_117, 0, x_113); +lean_ctor_set(x_117, 1, x_116); +return x_117; } 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_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* 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); 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; +x_120 = l_Lean_IR_formatFnBodyHead___closed__20; +x_131 = lean_unsigned_to_nat(1u); +x_132 = lean_nat_dec_eq(x_119, 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; lean_object* x_140; uint8_t x_141; lean_object* x_142; +x_133 = l_Nat_reprFast(x_119); +x_134 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_134, 0, x_133); +x_135 = l_Lean_IR_formatFnBodyHead___closed__22; +x_136 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_137 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_134); +x_138 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +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 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_140, 0, x_135); +lean_ctor_set(x_140, 1, x_139); +x_141 = 0; +x_142 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set_uint8(x_142, sizeof(void*)*1, x_141); +x_121 = x_142; +goto block_130; } else { -lean_object* x_136; -lean_dec(x_112); -x_136 = l_Lean_IR_formatFnBodyHead___closed__23; -x_114 = x_136; -goto block_123; +lean_object* x_143; +lean_dec(x_119); +x_143 = l_Lean_IR_formatFnBodyHead___closed__23; +x_121 = x_143; +goto block_130; } -block_123: +block_130: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_115 = lean_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_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); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; x_122 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_122, 0, x_117); +lean_ctor_set(x_122, 0, x_120); lean_ctor_set(x_122, 1, x_121); -return x_122; +x_123 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; +x_124 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +x_125 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_126 = l_Nat_reprFast(x_118); +x_127 = lean_string_append(x_125, x_126); +lean_dec_ref(x_126); +x_128 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_128, 0, x_127); +x_129 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_129, 0, x_124); +lean_ctor_set(x_129, 1, x_128); +return x_129; } } case 7: { -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_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_157; uint8_t x_158; +x_144 = lean_ctor_get(x_1, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_1, 1); +lean_inc(x_145); 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) -{ -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; +x_146 = l_Lean_IR_formatFnBodyHead___closed__25; +x_157 = lean_unsigned_to_nat(1u); +x_158 = lean_nat_dec_eq(x_145, x_157); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; +x_159 = l_Nat_reprFast(x_145); +x_160 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_160, 0, x_159); +x_161 = l_Lean_IR_formatFnBodyHead___closed__22; +x_162 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_163 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_160); +x_164 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +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 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_166, 0, x_161); +lean_ctor_set(x_166, 1, x_165); +x_167 = 0; +x_168 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set_uint8(x_168, sizeof(void*)*1, x_167); +x_147 = x_168; +goto block_156; } else { -lean_object* x_162; -lean_dec(x_138); -x_162 = l_Lean_IR_formatFnBodyHead___closed__23; -x_140 = x_162; -goto block_149; -} -block_149: -{ -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); +lean_object* x_169; +lean_dec(x_145); +x_169 = l_Lean_IR_formatFnBodyHead___closed__23; +x_147 = x_169; +goto block_156; +} +block_156: +{ +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_148 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_148, 0, x_143); +lean_ctor_set(x_148, 0, x_146); lean_ctor_set(x_148, 1, x_147); -return x_148; +x_149 = l_Lean_IR_formatArray___redArg___lam__0___closed__1; +x_150 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +x_151 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_152 = l_Nat_reprFast(x_144); +x_153 = lean_string_append(x_151, x_152); +lean_dec_ref(x_152); +x_154 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_154, 0, x_153); +x_155 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_155, 0, x_150); +lean_ctor_set(x_155, 1, x_154); +return x_155; } } 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); -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_170); -lean_ctor_set(x_1, 0, x_166); -return x_1; -} -else +uint8_t x_170; +x_170 = !lean_is_exclusive(x_1); +if (x_170 == 0) { 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_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_172 = lean_ctor_get(x_1, 1); +lean_dec(x_172); +x_173 = l_Lean_IR_formatFnBodyHead___closed__27; +x_174 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_175 = l_Nat_reprFast(x_171); +x_176 = lean_string_append(x_174, x_175); +lean_dec_ref(x_175); +x_177 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_177); +lean_ctor_set(x_1, 0, x_173); +return x_1; } -case 9: +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; -x_178 = lean_ctor_get(x_1, 1); +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_178 = lean_ctor_get(x_1, 0); lean_inc(x_178); -lean_dec_ref(x_1); -x_179 = l_Lean_IR_formatFnBodyHead___closed__29; +lean_dec(x_1); +x_179 = l_Lean_IR_formatFnBodyHead___closed__27; 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); @@ -3581,82 +3726,101 @@ 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; +return x_184; +} +} +case 9: +{ +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_185 = lean_ctor_get(x_1, 1); +lean_inc(x_185); +lean_dec_ref(x_1); +x_186 = l_Lean_IR_formatFnBodyHead___closed__29; +x_187 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_188 = l_Nat_reprFast(x_185); +x_189 = lean_string_append(x_187, x_188); +lean_dec_ref(x_188); +x_190 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_190, 0, x_189); +x_191 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_191, 0, x_186); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Lean_IR_formatFnBodyHead___closed__31; +x_193 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set(x_193, 1, x_192); +return x_193; } 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_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_194 = lean_ctor_get(x_1, 0); +lean_inc(x_194); 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_195 = l_Lean_IR_formatFnBodyHead___closed__33; +x_196 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_194); +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 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_198; +x_198 = !lean_is_exclusive(x_1); +if (x_198 == 0) +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_199 = lean_ctor_get(x_1, 0); +x_200 = lean_ctor_get(x_1, 1); +x_201 = l_Lean_IR_formatFnBodyHead___closed__35; +x_202 = l_Lean_IR_formatFnBodyHead___closed__4; +x_203 = l_Nat_reprFast(x_199); +x_204 = lean_string_append(x_202, x_203); +lean_dec_ref(x_203); +x_205 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_205, 0, x_204); 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_205); +lean_ctor_set(x_1, 0, x_201); +x_206 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_200); +lean_dec_ref(x_200); +x_207 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_207, 0, x_1); +lean_ctor_set(x_207, 1, x_206); +return x_207; } 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_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +x_208 = lean_ctor_get(x_1, 0); +x_209 = lean_ctor_get(x_1, 1); +lean_inc(x_209); +lean_inc(x_208); 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_210 = l_Lean_IR_formatFnBodyHead___closed__35; +x_211 = l_Lean_IR_formatFnBodyHead___closed__4; +x_212 = l_Nat_reprFast(x_208); +x_213 = lean_string_append(x_211, x_212); +lean_dec_ref(x_212); +x_214 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_214, 0, x_213); +x_215 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_215, 0, x_210); +lean_ctor_set(x_215, 1, x_214); +x_216 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_209); +lean_dec_ref(x_209); +x_217 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_217, 0, x_215); +lean_ctor_set(x_217, 1, x_216); +return x_217; } } default: { -lean_object* x_211; -x_211 = l_Lean_IR_formatFnBodyHead___closed__37; -return x_211; +lean_object* x_218; +x_218 = l_Lean_IR_formatFnBodyHead___closed__37; +return x_218; } } } @@ -3970,518 +4134,542 @@ 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; 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; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; x_95 = lean_ctor_get(x_2, 0); lean_inc(x_95); x_96 = lean_ctor_get(x_2, 1); lean_inc(x_96); 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_98 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); +x_99 = lean_ctor_get(x_2, 3); +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_114 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_114, 0, x_113); +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); +if (x_98 == 0) +{ +lean_object* x_127; +x_127 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__14; +x_111 = x_127; +goto block_126; +} +else +{ +lean_object* x_128; +x_128 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__15; +x_111 = x_128; +goto block_126; +} +block_126: +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_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_110); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_IR_formatFnBodyHead___closed__10; x_115 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); -x_116 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; -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_116 = l_Nat_reprFast(x_97); +x_117 = lean_string_append(x_101, x_116); +lean_dec_ref(x_116); +x_118 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_118, 0, x_117); x_119 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 0, x_115); 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_120 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; x_121 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_121, 0, x_119); lean_ctor_set(x_121, 1, x_120); -return x_121; +x_122 = lean_box(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 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_99); +x_125 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} } 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_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; +x_129 = lean_ctor_get(x_2, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_2, 1); +lean_inc(x_130); +x_131 = lean_ctor_get(x_2, 2); +lean_inc(x_131); +x_132 = lean_ctor_get(x_2, 3); +lean_inc(x_132); +x_133 = lean_ctor_get(x_2, 4); +lean_inc(x_133); +x_134 = lean_ctor_get(x_2, 5); +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_129); +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_130); +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__19; 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_131); +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 = l_Lean_IR_formatFnBodyHead___closed__18; +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___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_133); +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_Lean_IR_formatFnBodyHead___closed__3; +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_Nat_reprFast(x_132); +x_158 = lean_string_append(x_136, x_157); +lean_dec_ref(x_157); +x_159 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_159, 0, x_158); +x_160 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_160, 0, x_156); +lean_ctor_set(x_160, 1, x_159); +x_161 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_162 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +x_163 = lean_box(1); +x_164 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +x_165 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_134); +x_166 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_166, 0, x_164); +lean_ctor_set(x_166, 1, x_165); +return x_166; } 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_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_187; uint8_t x_188; +x_167 = lean_ctor_get(x_2, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_2, 1); +lean_inc(x_168); +x_169 = lean_ctor_get(x_2, 2); +lean_inc(x_169); 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_170 = l_Lean_IR_formatFnBodyHead___closed__20; +x_187 = lean_unsigned_to_nat(1u); +x_188 = lean_nat_dec_eq(x_168, x_187); +if (x_188 == 0) +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; +x_189 = l_Nat_reprFast(x_168); +x_190 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_190, 0, x_189); +x_191 = l_Lean_IR_formatFnBodyHead___closed__22; +x_192 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_193 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_190); +x_194 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_195 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_196, 0, x_191); +lean_ctor_set(x_196, 1, x_195); +x_197 = 0; +x_198 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set_uint8(x_198, sizeof(void*)*1, x_197); +x_171 = x_198; +goto block_186; } 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_199; +lean_dec(x_168); +x_199 = l_Lean_IR_formatFnBodyHead___closed__23; +x_171 = x_199; +goto block_186; } -block_179: +block_186: { -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); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; x_172 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_172, 0, x_167); +lean_ctor_set(x_172, 0, x_170); lean_ctor_set(x_172, 1, x_171); -x_173 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_173 = l_Lean_IR_formatArray___redArg___lam__0___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); -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_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_175 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_176 = l_Nat_reprFast(x_167); +x_177 = lean_string_append(x_175, x_176); +lean_dec_ref(x_176); +x_178 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_178, 0, x_177); +x_179 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_179, 0, x_174); +lean_ctor_set(x_179, 1, x_178); +x_180 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_181 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_180); +x_182 = lean_box(1); +x_183 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_183, 0, x_181); +lean_ctor_set(x_183, 1, x_182); +x_184 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_169); +x_185 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +return x_185; } } 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_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_220; uint8_t x_221; +x_200 = lean_ctor_get(x_2, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_2, 1); +lean_inc(x_201); +x_202 = lean_ctor_get(x_2, 2); +lean_inc(x_202); 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_203 = l_Lean_IR_formatFnBodyHead___closed__25; +x_220 = lean_unsigned_to_nat(1u); +x_221 = lean_nat_dec_eq(x_201, x_220); +if (x_221 == 0) +{ +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; uint8_t x_230; lean_object* x_231; +x_222 = l_Nat_reprFast(x_201); +x_223 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_223, 0, x_222); +x_224 = l_Lean_IR_formatFnBodyHead___closed__22; +x_225 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__1; +x_226 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_223); +x_227 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatCtorInfo___closed__3; +x_228 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_228, 0, x_226); +lean_ctor_set(x_228, 1, x_227); +x_229 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_229, 0, x_224); +lean_ctor_set(x_229, 1, x_228); +x_230 = 0; +x_231 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set_uint8(x_231, sizeof(void*)*1, x_230); +x_204 = x_231; +goto block_219; } else { -lean_object* x_225; -lean_dec(x_194); -x_225 = l_Lean_IR_formatFnBodyHead___closed__23; -x_197 = x_225; -goto block_212; +lean_object* x_232; +lean_dec(x_201); +x_232 = l_Lean_IR_formatFnBodyHead___closed__23; +x_204 = x_232; +goto block_219; } -block_212: +block_219: { -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); +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; x_205 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_205, 0, x_200); +lean_ctor_set(x_205, 0, x_203); lean_ctor_set(x_205, 1, x_204); -x_206 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_206 = l_Lean_IR_formatArray___redArg___lam__0___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); -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_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_208 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_209 = l_Nat_reprFast(x_200); +x_210 = lean_string_append(x_208, x_209); +lean_dec_ref(x_209); +x_211 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_211, 0, x_210); +x_212 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_212, 0, x_207); +lean_ctor_set(x_212, 1, x_211); +x_213 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_214 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +x_215 = lean_box(1); +x_216 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +x_217 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_202); +x_218 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +return x_218; } } 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_233; +x_233 = !lean_is_exclusive(x_2); +if (x_233 == 0) +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_234 = lean_ctor_get(x_2, 0); +x_235 = lean_ctor_get(x_2, 1); +x_236 = l_Lean_IR_formatFnBodyHead___closed__27; +x_237 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_238 = l_Nat_reprFast(x_234); +x_239 = lean_string_append(x_237, x_238); +lean_dec_ref(x_238); +x_240 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_240, 0, x_239); 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_240); +lean_ctor_set(x_2, 0, x_236); +x_241 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_242 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_242, 0, x_2); +lean_ctor_set(x_242, 1, x_241); +x_243 = lean_box(1); +x_244 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_244, 0, x_242); +lean_ctor_set(x_244, 1, x_243); +x_245 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_235); +x_246 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_246, 0, x_244); +lean_ctor_set(x_246, 1, x_245); +return x_246; } 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_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_247 = lean_ctor_get(x_2, 0); +x_248 = lean_ctor_get(x_2, 1); +lean_inc(x_248); +lean_inc(x_247); 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; +x_249 = l_Lean_IR_formatFnBodyHead___closed__27; +x_250 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_251 = l_Nat_reprFast(x_247); +x_252 = lean_string_append(x_250, x_251); +lean_dec_ref(x_251); +x_253 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_253, 0, x_252); +x_254 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_254, 0, x_249); +lean_ctor_set(x_254, 1, x_253); +x_255 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__1; +x_256 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_257 = lean_box(1); +x_258 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_258, 0, x_256); +lean_ctor_set(x_258, 1, x_257); +x_259 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop(x_1, x_248); +x_260 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +return x_260; } } 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_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; uint8_t x_279; +x_261 = lean_ctor_get(x_2, 1); +lean_inc(x_261); +x_262 = lean_ctor_get(x_2, 2); +lean_inc(x_262); +x_263 = lean_ctor_get(x_2, 3); +lean_inc_ref(x_263); 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_262 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_262, 0, x_257); -lean_ctor_set(x_262, 1, x_261); -x_263 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; -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(x_271); -lean_dec_ref(x_256); -lean_dec(x_1); +x_264 = l_Lean_IR_formatFnBodyHead___closed__29; +x_265 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg___closed__0; +x_266 = l_Nat_reprFast(x_261); +x_267 = lean_string_append(x_265, x_266); +lean_dec_ref(x_266); +x_268 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_268, 0, x_267); +x_269 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_269, 0, x_264); +lean_ctor_set(x_269, 1, x_268); +x_270 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatParam___closed__3; +x_271 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_271, 0, x_269); +lean_ctor_set(x_271, 1, x_270); +x_272 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType(x_262); 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, 0, x_271); +lean_ctor_set(x_273, 1, x_272); +x_274 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatFnBody_loop___closed__5; +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 = lean_box(0); +x_277 = lean_unsigned_to_nat(0u); +x_278 = lean_array_get_size(x_263); +x_279 = lean_nat_dec_lt(x_277, x_278); +if (x_279 == 0) +{ +lean_object* x_280; +lean_dec(x_278); +lean_dec_ref(x_263); +lean_dec(x_1); +x_280 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_280, 0, x_275); +lean_ctor_set(x_280, 1, x_276); +return x_280; } else { -uint8_t x_274; -x_274 = lean_nat_dec_le(x_271, x_271); -if (x_274 == 0) +uint8_t x_281; +x_281 = lean_nat_dec_le(x_278, x_278); +if (x_281 == 0) { -lean_object* x_275; -lean_dec(x_271); -lean_dec_ref(x_256); +lean_object* x_282; +lean_dec(x_278); +lean_dec_ref(x_263); 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_282 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_282, 0, x_275); +lean_ctor_set(x_282, 1, x_276); +return x_282; } 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); -lean_dec(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_283; size_t x_284; lean_object* x_285; lean_object* x_286; +x_283 = 0; +x_284 = lean_usize_of_nat(x_278); +lean_dec(x_278); +x_285 = 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_263, x_283, x_284, x_276); +lean_dec_ref(x_263); +x_286 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_286, 0, x_275); +lean_ctor_set(x_286, 1, x_285); +return x_286; } } } case 10: { -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_dec(x_1); -x_280 = lean_ctor_get(x_2, 0); -lean_inc(x_280); +x_287 = lean_ctor_get(x_2, 0); +lean_inc(x_287); lean_dec_ref(x_2); -x_281 = l_Lean_IR_formatFnBodyHead___closed__33; -x_282 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_280); -x_283 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_283, 0, x_281); -lean_ctor_set(x_283, 1, x_282); -return x_283; +x_288 = l_Lean_IR_formatFnBodyHead___closed__33; +x_289 = l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatArg(x_287); +x_290 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_290, 0, x_288); +lean_ctor_set(x_290, 1, x_289); +return x_290; } case 11: { -uint8_t x_284; +uint8_t x_291; lean_dec(x_1); -x_284 = !lean_is_exclusive(x_2); -if (x_284 == 0) -{ -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_285 = lean_ctor_get(x_2, 0); -x_286 = lean_ctor_get(x_2, 1); -x_287 = l_Lean_IR_formatFnBodyHead___closed__35; -x_288 = l_Lean_IR_formatFnBodyHead___closed__4; -x_289 = l_Nat_reprFast(x_285); -x_290 = lean_string_append(x_288, x_289); -lean_dec_ref(x_289); -x_291 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_291, 0, x_290); +x_291 = !lean_is_exclusive(x_2); +if (x_291 == 0) +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_292 = lean_ctor_get(x_2, 0); +x_293 = lean_ctor_get(x_2, 1); +x_294 = l_Lean_IR_formatFnBodyHead___closed__35; +x_295 = l_Lean_IR_formatFnBodyHead___closed__4; +x_296 = l_Nat_reprFast(x_292); +x_297 = lean_string_append(x_295, x_296); +lean_dec_ref(x_296); +x_298 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_298, 0, x_297); lean_ctor_set_tag(x_2, 5); -lean_ctor_set(x_2, 1, x_291); -lean_ctor_set(x_2, 0, x_287); -x_292 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_286); -lean_dec_ref(x_286); -x_293 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_293, 0, x_2); -lean_ctor_set(x_293, 1, x_292); -return x_293; +lean_ctor_set(x_2, 1, x_298); +lean_ctor_set(x_2, 0, x_294); +x_299 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_293); +lean_dec_ref(x_293); +x_300 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_300, 0, x_2); +lean_ctor_set(x_300, 1, x_299); +return x_300; } else { -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; -x_294 = lean_ctor_get(x_2, 0); -x_295 = lean_ctor_get(x_2, 1); -lean_inc(x_295); -lean_inc(x_294); +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; +x_301 = lean_ctor_get(x_2, 0); +x_302 = lean_ctor_get(x_2, 1); +lean_inc(x_302); +lean_inc(x_301); lean_dec(x_2); -x_296 = l_Lean_IR_formatFnBodyHead___closed__35; -x_297 = l_Lean_IR_formatFnBodyHead___closed__4; -x_298 = l_Nat_reprFast(x_294); -x_299 = lean_string_append(x_297, x_298); -lean_dec_ref(x_298); -x_300 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_300, 0, x_299); -x_301 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_301, 0, x_296); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_295); -lean_dec_ref(x_295); -x_303 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_303, 0, x_301); -lean_ctor_set(x_303, 1, x_302); -return x_303; +x_303 = l_Lean_IR_formatFnBodyHead___closed__35; +x_304 = l_Lean_IR_formatFnBodyHead___closed__4; +x_305 = l_Nat_reprFast(x_301); +x_306 = lean_string_append(x_304, x_305); +lean_dec_ref(x_305); +x_307 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_307, 0, x_306); +x_308 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_308, 0, x_303); +lean_ctor_set(x_308, 1, x_307); +x_309 = l_Lean_IR_formatArray___at___00__private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr_spec__0(x_302); +lean_dec_ref(x_302); +x_310 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_310, 0, x_308); +lean_ctor_set(x_310, 1, x_309); +return x_310; } } default: { -lean_object* x_304; +lean_object* x_311; lean_dec(x_1); -x_304 = l_Lean_IR_formatFnBodyHead___closed__37; -return x_304; +x_311 = l_Lean_IR_formatFnBodyHead___closed__37; +return x_311; } } } @@ -4839,6 +5027,10 @@ l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__26 = _init_l lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__26); l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27(); lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__27); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__28); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatExpr___closed__29); l_Lean_IR_instToFormatExpr___closed__0 = _init_l_Lean_IR_instToFormatExpr___closed__0(); lean_mark_persistent(l_Lean_IR_instToFormatExpr___closed__0); l_Lean_IR_instToFormatExpr = _init_l_Lean_IR_instToFormatExpr(); @@ -4911,6 +5103,26 @@ l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__30 = _init 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__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___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__37); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__38 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__38(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__38); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__39); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__40); +l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41 = _init_l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41(); +lean_mark_persistent(l___private_Lean_Compiler_IR_Format_0__Lean_IR_formatIRType___closed__41); 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(); diff --git a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c index 480ee581c800..3131265edb49 100644 --- a/stage0/stdlib/Lean/Compiler/IR/FreeVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/FreeVars.c @@ -1484,9 +1484,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 { @@ -1505,7 +1505,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; } @@ -1531,8 +1531,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; } @@ -1540,8 +1540,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; } @@ -1957,9 +1957,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 { @@ -1978,7 +1978,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; } @@ -2004,8 +2004,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; } @@ -2013,8 +2013,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; } diff --git a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c index 7630022a28cc..12bfc8f60cd6 100644 --- a/stage0/stdlib/Lean/Compiler/IR/LiveVars.c +++ b/stage0/stdlib/Lean/Compiler/IR/LiveVars.c @@ -1472,9 +1472,9 @@ 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); +x_141 = lean_nat_add(x_138, x_140); lean_dec(x_140); -lean_dec(x_139); +lean_dec(x_138); if (lean_is_scalar(x_135)) { x_142 = lean_alloc_ctor(0, 5, 0); } else { @@ -1493,7 +1493,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_138); +lean_ctor_set(x_143, 3, x_139); lean_ctor_set(x_143, 4, x_142); return x_143; } @@ -1519,8 +1519,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_148; -x_139 = x_149; +x_138 = x_149; +x_139 = x_148; x_140 = x_150; goto block_144; } @@ -1528,8 +1528,8 @@ else { lean_object* x_151; x_151 = lean_unsigned_to_nat(0u); -x_138 = x_148; -x_139 = x_149; +x_138 = x_149; +x_139 = x_148; x_140 = x_151; goto block_144; } @@ -3216,9 +3216,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); @@ -3245,7 +3245,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); @@ -3259,7 +3259,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; } @@ -3286,8 +3286,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; } @@ -3295,8 +3295,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; } @@ -3382,9 +3382,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); @@ -3415,7 +3415,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; } @@ -3441,8 +3441,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; } @@ -3450,8 +3450,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; } @@ -3920,9 +3920,9 @@ goto block_449; block_442: { lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_439 = lean_nat_add(x_436, x_438); +x_439 = lean_nat_add(x_437, x_438); lean_dec(x_438); -lean_dec(x_436); +lean_dec(x_437); if (lean_is_scalar(x_433)) { x_440 = lean_alloc_ctor(0, 5, 0); } else { @@ -3941,7 +3941,7 @@ if (lean_is_scalar(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_437); +lean_ctor_set(x_441, 3, x_436); lean_ctor_set(x_441, 4, x_440); return x_441; } @@ -3967,8 +3967,8 @@ 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_446; -x_437 = x_445; +x_436 = x_445; +x_437 = x_446; x_438 = x_447; goto block_442; } @@ -3976,8 +3976,8 @@ else { lean_object* x_448; x_448 = lean_unsigned_to_nat(0u); -x_436 = x_446; -x_437 = x_445; +x_436 = x_445; +x_437 = x_446; x_438 = x_448; goto block_442; } diff --git a/stage0/stdlib/Lean/Compiler/IR/NormIds.c b/stage0/stdlib/Lean/Compiler/IR/NormIds.c index f37c026ca625..544c7df8a8b0 100644 --- a/stage0/stdlib/Lean/Compiler/IR/NormIds.c +++ b/stage0/stdlib/Lean/Compiler/IR/NormIds.c @@ -2208,192 +2208,194 @@ return x_1; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; x_39 = lean_ctor_get(x_1, 0); -x_40 = lean_ctor_get(x_1, 1); -lean_inc(x_40); +x_40 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +x_41 = lean_ctor_get(x_1, 1); +lean_inc(x_41); 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_42 = l_Lean_IR_NormalizeIds_normIndex(x_41, x_2); +lean_dec(x_41); +x_43 = lean_alloc_ctor(4, 2, 1); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*2, x_40); +return x_43; } } case 5: { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_1); -if (x_43 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_1); +if (x_44 == 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_45; lean_object* x_46; +x_45 = lean_ctor_get(x_1, 2); +x_46 = l_Lean_IR_NormalizeIds_normIndex(x_45, x_2); +lean_dec(x_45); +lean_ctor_set(x_1, 2, x_46); 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_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_1, 0); +x_48 = lean_ctor_get(x_1, 1); +x_49 = lean_ctor_get(x_1, 2); +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_50 = l_Lean_IR_NormalizeIds_normIndex(x_49, x_2); +lean_dec(x_49); +x_51 = lean_alloc_ctor(5, 3, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_48); +lean_ctor_set(x_51, 2, x_50); +return x_51; } } case 6: { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_1); -if (x_51 == 0) +uint8_t x_52; +x_52 = !lean_is_exclusive(x_1); +if (x_52 == 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_53; lean_object* x_54; +x_53 = lean_ctor_get(x_1, 1); +x_54 = l_Lean_IR_NormalizeIds_normArgs(x_53, x_2); +lean_ctor_set(x_1, 1, x_54); 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_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_1, 0); +x_56 = lean_ctor_get(x_1, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); 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_57 = l_Lean_IR_NormalizeIds_normArgs(x_56, x_2); +x_58 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_57); +return x_58; } } case 7: { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_1); -if (x_58 == 0) +uint8_t x_59; +x_59 = !lean_is_exclusive(x_1); +if (x_59 == 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_60; lean_object* x_61; +x_60 = lean_ctor_get(x_1, 1); +x_61 = l_Lean_IR_NormalizeIds_normArgs(x_60, x_2); +lean_ctor_set(x_1, 1, x_61); 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_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_1, 0); +x_63 = lean_ctor_get(x_1, 1); +lean_inc(x_63); lean_inc(x_62); -lean_inc(x_61); 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_64 = l_Lean_IR_NormalizeIds_normArgs(x_63, x_2); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } case 8: { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_1); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_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); +uint8_t x_66; +x_66 = !lean_is_exclusive(x_1); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_1, 0); +x_68 = lean_ctor_get(x_1, 1); +x_69 = l_Lean_IR_NormalizeIds_normIndex(x_67, x_2); +lean_dec(x_67); +x_70 = l_Lean_IR_NormalizeIds_normArgs(x_68, x_2); +lean_ctor_set(x_1, 1, x_70); +lean_ctor_set(x_1, 0, x_69); return x_1; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_1, 0); -x_71 = lean_ctor_get(x_1, 1); +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_1, 0); +x_72 = lean_ctor_get(x_1, 1); +lean_inc(x_72); lean_inc(x_71); -lean_inc(x_70); 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_73 = l_Lean_IR_NormalizeIds_normIndex(x_71, x_2); +lean_dec(x_71); +x_74 = l_Lean_IR_NormalizeIds_normArgs(x_72, x_2); +x_75 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } case 9: { -uint8_t x_75; -x_75 = !lean_is_exclusive(x_1); -if (x_75 == 0) +uint8_t x_76; +x_76 = !lean_is_exclusive(x_1); +if (x_76 == 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_77; lean_object* x_78; +x_77 = lean_ctor_get(x_1, 1); +x_78 = l_Lean_IR_NormalizeIds_normIndex(x_77, x_2); +lean_dec(x_77); +lean_ctor_set(x_1, 1, x_78); 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_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_79 = lean_ctor_get(x_1, 0); +x_80 = lean_ctor_get(x_1, 1); +lean_inc(x_80); lean_inc(x_79); -lean_inc(x_78); 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_81 = l_Lean_IR_NormalizeIds_normIndex(x_80, x_2); +lean_dec(x_80); +x_82 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_82, 0, x_79); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } case 10: { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) +uint8_t x_83; +x_83 = !lean_is_exclusive(x_1); +if (x_83 == 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_84; lean_object* x_85; +x_84 = lean_ctor_get(x_1, 0); +x_85 = l_Lean_IR_NormalizeIds_normIndex(x_84, x_2); +lean_dec(x_84); +lean_ctor_set(x_1, 0, x_85); 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_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_1, 0); +lean_inc(x_86); 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_87 = l_Lean_IR_NormalizeIds_normIndex(x_86, x_2); +lean_dec(x_86); +x_88 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_88, 0, x_87); +return x_88; } } case 11: @@ -2402,28 +2404,28 @@ return x_1; } default: { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_1); -if (x_88 == 0) +uint8_t x_89; +x_89 = !lean_is_exclusive(x_1); +if (x_89 == 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_90; lean_object* x_91; +x_90 = lean_ctor_get(x_1, 0); +x_91 = l_Lean_IR_NormalizeIds_normIndex(x_90, x_2); +lean_dec(x_90); +lean_ctor_set(x_1, 0, x_91); 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_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_1, 0); +lean_inc(x_92); 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_93 = l_Lean_IR_NormalizeIds_normIndex(x_92, x_2); +lean_dec(x_92); +x_94 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_94, 0, x_93); +return x_94; } } } @@ -3648,581 +3650,583 @@ return x_130; } else { -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_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; x_131 = lean_ctor_get(x_1, 0); x_132 = lean_ctor_get(x_1, 1); x_133 = lean_ctor_get(x_1, 2); -x_134 = lean_ctor_get(x_1, 3); -lean_inc(x_134); +x_134 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_135 = lean_ctor_get(x_1, 3); +lean_inc(x_135); lean_inc(x_133); lean_inc(x_132); lean_inc(x_131); lean_dec(x_1); lean_inc(x_2); -x_135 = l_Lean_IR_NormalizeIds_normFnBody(x_134, x_2, x_3); -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_135, 1); +x_136 = l_Lean_IR_NormalizeIds_normFnBody(x_135, x_2, x_3); +x_137 = lean_ctor_get(x_136, 0); lean_inc(x_137); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_138 = x_135; +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_139 = x_136; } else { - lean_dec_ref(x_135); - x_138 = lean_box(0); + lean_dec_ref(x_136); + x_139 = lean_box(0); } -x_139 = l_Lean_IR_NormalizeIds_normIndex(x_131, x_2); +x_140 = l_Lean_IR_NormalizeIds_normIndex(x_131, x_2); lean_dec(x_131); -x_140 = l_Lean_IR_NormalizeIds_normIndex(x_133, x_2); +x_141 = l_Lean_IR_NormalizeIds_normIndex(x_133, x_2); lean_dec(x_2); lean_dec(x_133); -x_141 = lean_alloc_ctor(4, 4, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_132); -lean_ctor_set(x_141, 2, x_140); -lean_ctor_set(x_141, 3, x_136); -if (lean_is_scalar(x_138)) { - x_142 = lean_alloc_ctor(0, 2, 0); +x_142 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_132); +lean_ctor_set(x_142, 2, x_141); +lean_ctor_set(x_142, 3, x_137); +lean_ctor_set_uint8(x_142, sizeof(void*)*4, x_134); +if (lean_is_scalar(x_139)) { + x_143 = lean_alloc_ctor(0, 2, 0); } else { - x_142 = x_138; + x_143 = x_139; } -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_137); -return x_142; +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_138); +return x_143; } } case 5: { -uint8_t x_143; -x_143 = !lean_is_exclusive(x_1); -if (x_143 == 0) +uint8_t x_144; +x_144 = !lean_is_exclusive(x_1); +if (x_144 == 0) { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; -x_144 = lean_ctor_get(x_1, 0); -x_145 = lean_ctor_get(x_1, 3); -x_146 = lean_ctor_get(x_1, 5); +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; +x_145 = lean_ctor_get(x_1, 0); +x_146 = lean_ctor_get(x_1, 3); +x_147 = lean_ctor_get(x_1, 5); lean_inc(x_2); -x_147 = l_Lean_IR_NormalizeIds_normFnBody(x_146, x_2, x_3); -x_148 = !lean_is_exclusive(x_147); -if (x_148 == 0) -{ -lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_147, 0); -x_150 = l_Lean_IR_NormalizeIds_normIndex(x_144, x_2); -lean_dec(x_144); +x_148 = l_Lean_IR_NormalizeIds_normFnBody(x_147, x_2, x_3); +x_149 = !lean_is_exclusive(x_148); +if (x_149 == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_148, 0); x_151 = l_Lean_IR_NormalizeIds_normIndex(x_145, x_2); -lean_dec(x_2); lean_dec(x_145); -lean_ctor_set(x_1, 5, x_149); -lean_ctor_set(x_1, 3, x_151); -lean_ctor_set(x_1, 0, x_150); -lean_ctor_set(x_147, 0, x_1); -return x_147; +x_152 = l_Lean_IR_NormalizeIds_normIndex(x_146, x_2); +lean_dec(x_2); +lean_dec(x_146); +lean_ctor_set(x_1, 5, x_150); +lean_ctor_set(x_1, 3, x_152); +lean_ctor_set(x_1, 0, x_151); +lean_ctor_set(x_148, 0, x_1); +return x_148; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_152 = lean_ctor_get(x_147, 0); -x_153 = lean_ctor_get(x_147, 1); +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_153 = lean_ctor_get(x_148, 0); +x_154 = lean_ctor_get(x_148, 1); +lean_inc(x_154); lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_147); -x_154 = l_Lean_IR_NormalizeIds_normIndex(x_144, x_2); -lean_dec(x_144); +lean_dec(x_148); x_155 = l_Lean_IR_NormalizeIds_normIndex(x_145, x_2); -lean_dec(x_2); lean_dec(x_145); -lean_ctor_set(x_1, 5, x_152); -lean_ctor_set(x_1, 3, x_155); -lean_ctor_set(x_1, 0, x_154); -x_156 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_156, 0, x_1); -lean_ctor_set(x_156, 1, x_153); -return x_156; +x_156 = l_Lean_IR_NormalizeIds_normIndex(x_146, x_2); +lean_dec(x_2); +lean_dec(x_146); +lean_ctor_set(x_1, 5, x_153); +lean_ctor_set(x_1, 3, x_156); +lean_ctor_set(x_1, 0, x_155); +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_1); +lean_ctor_set(x_157, 1, x_154); +return x_157; } } 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; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_157 = lean_ctor_get(x_1, 0); -x_158 = lean_ctor_get(x_1, 1); -x_159 = lean_ctor_get(x_1, 2); -x_160 = lean_ctor_get(x_1, 3); -x_161 = lean_ctor_get(x_1, 4); -x_162 = lean_ctor_get(x_1, 5); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_158 = lean_ctor_get(x_1, 0); +x_159 = lean_ctor_get(x_1, 1); +x_160 = lean_ctor_get(x_1, 2); +x_161 = lean_ctor_get(x_1, 3); +x_162 = lean_ctor_get(x_1, 4); +x_163 = lean_ctor_get(x_1, 5); +lean_inc(x_163); lean_inc(x_162); lean_inc(x_161); lean_inc(x_160); lean_inc(x_159); lean_inc(x_158); -lean_inc(x_157); lean_dec(x_1); lean_inc(x_2); -x_163 = l_Lean_IR_NormalizeIds_normFnBody(x_162, x_2, x_3); -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); +x_164 = l_Lean_IR_NormalizeIds_normFnBody(x_163, x_2, x_3); +x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_166 = x_163; +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_163); - x_166 = lean_box(0); + lean_dec_ref(x_164); + x_167 = lean_box(0); } -x_167 = l_Lean_IR_NormalizeIds_normIndex(x_157, x_2); -lean_dec(x_157); -x_168 = l_Lean_IR_NormalizeIds_normIndex(x_160, x_2); +x_168 = l_Lean_IR_NormalizeIds_normIndex(x_158, x_2); +lean_dec(x_158); +x_169 = l_Lean_IR_NormalizeIds_normIndex(x_161, x_2); lean_dec(x_2); -lean_dec(x_160); -x_169 = lean_alloc_ctor(5, 6, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_158); -lean_ctor_set(x_169, 2, x_159); -lean_ctor_set(x_169, 3, x_168); -lean_ctor_set(x_169, 4, x_161); -lean_ctor_set(x_169, 5, x_164); -if (lean_is_scalar(x_166)) { - x_170 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_161); +x_170 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_159); +lean_ctor_set(x_170, 2, x_160); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_162); +lean_ctor_set(x_170, 5, x_165); +if (lean_is_scalar(x_167)) { + x_171 = lean_alloc_ctor(0, 2, 0); } else { - x_170 = x_166; + x_171 = x_167; } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_165); -return x_170; +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_166); +return x_171; } } case 6: { -uint8_t x_171; -x_171 = !lean_is_exclusive(x_1); -if (x_171 == 0) +uint8_t x_172; +x_172 = !lean_is_exclusive(x_1); +if (x_172 == 0) { -lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; -x_172 = lean_ctor_get(x_1, 0); -x_173 = lean_ctor_get(x_1, 2); +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; +x_173 = lean_ctor_get(x_1, 0); +x_174 = lean_ctor_get(x_1, 2); lean_inc(x_2); -x_174 = l_Lean_IR_NormalizeIds_normFnBody(x_173, x_2, x_3); -x_175 = !lean_is_exclusive(x_174); -if (x_175 == 0) +x_175 = l_Lean_IR_NormalizeIds_normFnBody(x_174, x_2, x_3); +x_176 = !lean_is_exclusive(x_175); +if (x_176 == 0) { -lean_object* x_176; lean_object* x_177; -x_176 = lean_ctor_get(x_174, 0); -x_177 = l_Lean_IR_NormalizeIds_normIndex(x_172, x_2); +lean_object* x_177; lean_object* x_178; +x_177 = lean_ctor_get(x_175, 0); +x_178 = l_Lean_IR_NormalizeIds_normIndex(x_173, x_2); lean_dec(x_2); -lean_dec(x_172); -lean_ctor_set(x_1, 2, x_176); -lean_ctor_set(x_1, 0, x_177); -lean_ctor_set(x_174, 0, x_1); -return x_174; +lean_dec(x_173); +lean_ctor_set(x_1, 2, x_177); +lean_ctor_set(x_1, 0, x_178); +lean_ctor_set(x_175, 0, x_1); +return x_175; } else { -lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_178 = lean_ctor_get(x_174, 0); -x_179 = lean_ctor_get(x_174, 1); +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_179 = lean_ctor_get(x_175, 0); +x_180 = lean_ctor_get(x_175, 1); +lean_inc(x_180); lean_inc(x_179); -lean_inc(x_178); -lean_dec(x_174); -x_180 = l_Lean_IR_NormalizeIds_normIndex(x_172, x_2); +lean_dec(x_175); +x_181 = l_Lean_IR_NormalizeIds_normIndex(x_173, x_2); lean_dec(x_2); -lean_dec(x_172); -lean_ctor_set(x_1, 2, x_178); -lean_ctor_set(x_1, 0, x_180); -x_181 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_181, 0, x_1); -lean_ctor_set(x_181, 1, x_179); -return x_181; +lean_dec(x_173); +lean_ctor_set(x_1, 2, x_179); +lean_ctor_set(x_1, 0, x_181); +x_182 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_182, 0, x_1); +lean_ctor_set(x_182, 1, x_180); +return x_182; } } else { -lean_object* x_182; lean_object* x_183; uint8_t x_184; uint8_t x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_182 = lean_ctor_get(x_1, 0); -x_183 = lean_ctor_get(x_1, 1); -x_184 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_185 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_186 = lean_ctor_get(x_1, 2); -lean_inc(x_186); +lean_object* x_183; lean_object* x_184; uint8_t x_185; uint8_t x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_183 = lean_ctor_get(x_1, 0); +x_184 = lean_ctor_get(x_1, 1); +x_185 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_186 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_187 = lean_ctor_get(x_1, 2); +lean_inc(x_187); +lean_inc(x_184); lean_inc(x_183); -lean_inc(x_182); lean_dec(x_1); lean_inc(x_2); -x_187 = l_Lean_IR_NormalizeIds_normFnBody(x_186, x_2, x_3); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); +x_188 = l_Lean_IR_NormalizeIds_normFnBody(x_187, x_2, x_3); +x_189 = lean_ctor_get(x_188, 0); lean_inc(x_189); -if (lean_is_exclusive(x_187)) { - lean_ctor_release(x_187, 0); - lean_ctor_release(x_187, 1); - x_190 = x_187; +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; } else { - lean_dec_ref(x_187); - x_190 = lean_box(0); + lean_dec_ref(x_188); + x_191 = lean_box(0); } -x_191 = l_Lean_IR_NormalizeIds_normIndex(x_182, x_2); +x_192 = l_Lean_IR_NormalizeIds_normIndex(x_183, x_2); lean_dec(x_2); -lean_dec(x_182); -x_192 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_183); -lean_ctor_set(x_192, 2, x_188); -lean_ctor_set_uint8(x_192, sizeof(void*)*3, x_184); -lean_ctor_set_uint8(x_192, sizeof(void*)*3 + 1, x_185); -if (lean_is_scalar(x_190)) { - x_193 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_183); +x_193 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_184); +lean_ctor_set(x_193, 2, x_189); +lean_ctor_set_uint8(x_193, sizeof(void*)*3, x_185); +lean_ctor_set_uint8(x_193, sizeof(void*)*3 + 1, x_186); +if (lean_is_scalar(x_191)) { + x_194 = lean_alloc_ctor(0, 2, 0); } else { - x_193 = x_190; + x_194 = x_191; } -lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_189); -return x_193; +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_190); +return x_194; } } case 7: { -uint8_t x_194; -x_194 = !lean_is_exclusive(x_1); -if (x_194 == 0) +uint8_t x_195; +x_195 = !lean_is_exclusive(x_1); +if (x_195 == 0) { -lean_object* x_195; lean_object* x_196; lean_object* x_197; uint8_t x_198; -x_195 = lean_ctor_get(x_1, 0); -x_196 = lean_ctor_get(x_1, 2); +lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; +x_196 = lean_ctor_get(x_1, 0); +x_197 = lean_ctor_get(x_1, 2); lean_inc(x_2); -x_197 = l_Lean_IR_NormalizeIds_normFnBody(x_196, x_2, x_3); -x_198 = !lean_is_exclusive(x_197); -if (x_198 == 0) +x_198 = l_Lean_IR_NormalizeIds_normFnBody(x_197, x_2, x_3); +x_199 = !lean_is_exclusive(x_198); +if (x_199 == 0) { -lean_object* x_199; lean_object* x_200; -x_199 = lean_ctor_get(x_197, 0); -x_200 = l_Lean_IR_NormalizeIds_normIndex(x_195, x_2); +lean_object* x_200; lean_object* x_201; +x_200 = lean_ctor_get(x_198, 0); +x_201 = l_Lean_IR_NormalizeIds_normIndex(x_196, x_2); lean_dec(x_2); -lean_dec(x_195); -lean_ctor_set(x_1, 2, x_199); -lean_ctor_set(x_1, 0, x_200); -lean_ctor_set(x_197, 0, x_1); -return x_197; +lean_dec(x_196); +lean_ctor_set(x_1, 2, x_200); +lean_ctor_set(x_1, 0, x_201); +lean_ctor_set(x_198, 0, x_1); +return x_198; } else { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_201 = lean_ctor_get(x_197, 0); -x_202 = lean_ctor_get(x_197, 1); +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_202 = lean_ctor_get(x_198, 0); +x_203 = lean_ctor_get(x_198, 1); +lean_inc(x_203); lean_inc(x_202); -lean_inc(x_201); -lean_dec(x_197); -x_203 = l_Lean_IR_NormalizeIds_normIndex(x_195, x_2); +lean_dec(x_198); +x_204 = l_Lean_IR_NormalizeIds_normIndex(x_196, x_2); lean_dec(x_2); -lean_dec(x_195); -lean_ctor_set(x_1, 2, x_201); -lean_ctor_set(x_1, 0, x_203); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_1); -lean_ctor_set(x_204, 1, x_202); -return x_204; +lean_dec(x_196); +lean_ctor_set(x_1, 2, x_202); +lean_ctor_set(x_1, 0, x_204); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_1); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } else { -lean_object* x_205; lean_object* x_206; uint8_t x_207; uint8_t x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_205 = lean_ctor_get(x_1, 0); -x_206 = lean_ctor_get(x_1, 1); -x_207 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_208 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); -x_209 = lean_ctor_get(x_1, 2); -lean_inc(x_209); +lean_object* x_206; lean_object* x_207; uint8_t 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; +x_206 = lean_ctor_get(x_1, 0); +x_207 = lean_ctor_get(x_1, 1); +x_208 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_209 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_210 = lean_ctor_get(x_1, 2); +lean_inc(x_210); +lean_inc(x_207); lean_inc(x_206); -lean_inc(x_205); lean_dec(x_1); lean_inc(x_2); -x_210 = l_Lean_IR_NormalizeIds_normFnBody(x_209, x_2, x_3); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); +x_211 = l_Lean_IR_NormalizeIds_normFnBody(x_210, x_2, x_3); +x_212 = lean_ctor_get(x_211, 0); lean_inc(x_212); -if (lean_is_exclusive(x_210)) { - lean_ctor_release(x_210, 0); - lean_ctor_release(x_210, 1); - x_213 = x_210; +x_213 = lean_ctor_get(x_211, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_211)) { + lean_ctor_release(x_211, 0); + lean_ctor_release(x_211, 1); + x_214 = x_211; } else { - lean_dec_ref(x_210); - x_213 = lean_box(0); + lean_dec_ref(x_211); + x_214 = lean_box(0); } -x_214 = l_Lean_IR_NormalizeIds_normIndex(x_205, x_2); +x_215 = l_Lean_IR_NormalizeIds_normIndex(x_206, x_2); lean_dec(x_2); -lean_dec(x_205); -x_215 = lean_alloc_ctor(7, 3, 2); -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_206); -lean_ctor_set(x_215, 2, x_211); -lean_ctor_set_uint8(x_215, sizeof(void*)*3, x_207); -lean_ctor_set_uint8(x_215, sizeof(void*)*3 + 1, x_208); -if (lean_is_scalar(x_213)) { - x_216 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_206); +x_216 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_207); +lean_ctor_set(x_216, 2, x_212); +lean_ctor_set_uint8(x_216, sizeof(void*)*3, x_208); +lean_ctor_set_uint8(x_216, sizeof(void*)*3 + 1, x_209); +if (lean_is_scalar(x_214)) { + x_217 = lean_alloc_ctor(0, 2, 0); } else { - x_216 = x_213; + x_217 = x_214; } -lean_ctor_set(x_216, 0, x_215); -lean_ctor_set(x_216, 1, x_212); -return x_216; +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_213); +return x_217; } } case 8: { -uint8_t x_217; -x_217 = !lean_is_exclusive(x_1); -if (x_217 == 0) +uint8_t x_218; +x_218 = !lean_is_exclusive(x_1); +if (x_218 == 0) { -lean_object* x_218; lean_object* x_219; lean_object* x_220; uint8_t x_221; -x_218 = lean_ctor_get(x_1, 0); -x_219 = lean_ctor_get(x_1, 1); +lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; +x_219 = lean_ctor_get(x_1, 0); +x_220 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_220 = l_Lean_IR_NormalizeIds_normFnBody(x_219, x_2, x_3); -x_221 = !lean_is_exclusive(x_220); -if (x_221 == 0) +x_221 = l_Lean_IR_NormalizeIds_normFnBody(x_220, x_2, x_3); +x_222 = !lean_is_exclusive(x_221); +if (x_222 == 0) { -lean_object* x_222; lean_object* x_223; -x_222 = lean_ctor_get(x_220, 0); -x_223 = l_Lean_IR_NormalizeIds_normIndex(x_218, x_2); +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_221, 0); +x_224 = l_Lean_IR_NormalizeIds_normIndex(x_219, x_2); lean_dec(x_2); -lean_dec(x_218); -lean_ctor_set(x_1, 1, x_222); -lean_ctor_set(x_1, 0, x_223); -lean_ctor_set(x_220, 0, x_1); -return x_220; +lean_dec(x_219); +lean_ctor_set(x_1, 1, x_223); +lean_ctor_set(x_1, 0, x_224); +lean_ctor_set(x_221, 0, x_1); +return x_221; } else { -lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_224 = lean_ctor_get(x_220, 0); -x_225 = lean_ctor_get(x_220, 1); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_225 = lean_ctor_get(x_221, 0); +x_226 = lean_ctor_get(x_221, 1); +lean_inc(x_226); lean_inc(x_225); -lean_inc(x_224); -lean_dec(x_220); -x_226 = l_Lean_IR_NormalizeIds_normIndex(x_218, x_2); +lean_dec(x_221); +x_227 = l_Lean_IR_NormalizeIds_normIndex(x_219, x_2); lean_dec(x_2); -lean_dec(x_218); -lean_ctor_set(x_1, 1, x_224); -lean_ctor_set(x_1, 0, x_226); -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_1); -lean_ctor_set(x_227, 1, x_225); -return x_227; +lean_dec(x_219); +lean_ctor_set(x_1, 1, x_225); +lean_ctor_set(x_1, 0, x_227); +x_228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_228, 0, x_1); +lean_ctor_set(x_228, 1, x_226); +return x_228; } } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_228 = lean_ctor_get(x_1, 0); -x_229 = lean_ctor_get(x_1, 1); +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_229 = lean_ctor_get(x_1, 0); +x_230 = lean_ctor_get(x_1, 1); +lean_inc(x_230); lean_inc(x_229); -lean_inc(x_228); lean_dec(x_1); lean_inc(x_2); -x_230 = l_Lean_IR_NormalizeIds_normFnBody(x_229, x_2, x_3); -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_ctor_get(x_230, 1); +x_231 = l_Lean_IR_NormalizeIds_normFnBody(x_230, x_2, x_3); +x_232 = lean_ctor_get(x_231, 0); lean_inc(x_232); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_233 = x_230; +x_233 = lean_ctor_get(x_231, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_234 = x_231; } else { - lean_dec_ref(x_230); - x_233 = lean_box(0); + lean_dec_ref(x_231); + x_234 = lean_box(0); } -x_234 = l_Lean_IR_NormalizeIds_normIndex(x_228, x_2); +x_235 = l_Lean_IR_NormalizeIds_normIndex(x_229, x_2); lean_dec(x_2); -lean_dec(x_228); -x_235 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_231); -if (lean_is_scalar(x_233)) { - x_236 = lean_alloc_ctor(0, 2, 0); -} else { - x_236 = x_233; -} +lean_dec(x_229); +x_236 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_236, 0, x_235); lean_ctor_set(x_236, 1, x_232); -return x_236; +if (lean_is_scalar(x_234)) { + x_237 = lean_alloc_ctor(0, 2, 0); +} else { + x_237 = x_234; +} +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_233); +return x_237; } } case 9: { -uint8_t x_237; -x_237 = !lean_is_exclusive(x_1); -if (x_237 == 0) +uint8_t x_238; +x_238 = !lean_is_exclusive(x_1); +if (x_238 == 0) { -lean_object* x_238; lean_object* x_239; size_t x_240; size_t x_241; lean_object* x_242; uint8_t x_243; -x_238 = lean_ctor_get(x_1, 1); -x_239 = lean_ctor_get(x_1, 3); -x_240 = lean_array_size(x_239); -x_241 = 0; +lean_object* x_239; lean_object* x_240; size_t x_241; size_t x_242; lean_object* x_243; uint8_t x_244; +x_239 = lean_ctor_get(x_1, 1); +x_240 = lean_ctor_get(x_1, 3); +x_241 = lean_array_size(x_240); +x_242 = 0; lean_inc(x_2); -x_242 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__3(x_240, x_241, x_239, x_2, x_3); -x_243 = !lean_is_exclusive(x_242); -if (x_243 == 0) +x_243 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__3(x_241, x_242, x_240, x_2, x_3); +x_244 = !lean_is_exclusive(x_243); +if (x_244 == 0) { -lean_object* x_244; lean_object* x_245; -x_244 = lean_ctor_get(x_242, 0); -x_245 = l_Lean_IR_NormalizeIds_normIndex(x_238, x_2); +lean_object* x_245; lean_object* x_246; +x_245 = lean_ctor_get(x_243, 0); +x_246 = l_Lean_IR_NormalizeIds_normIndex(x_239, x_2); lean_dec(x_2); -lean_dec(x_238); -lean_ctor_set(x_1, 3, x_244); -lean_ctor_set(x_1, 1, x_245); -lean_ctor_set(x_242, 0, x_1); -return x_242; +lean_dec(x_239); +lean_ctor_set(x_1, 3, x_245); +lean_ctor_set(x_1, 1, x_246); +lean_ctor_set(x_243, 0, x_1); +return x_243; } else { -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_246 = lean_ctor_get(x_242, 0); -x_247 = lean_ctor_get(x_242, 1); +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_247 = lean_ctor_get(x_243, 0); +x_248 = lean_ctor_get(x_243, 1); +lean_inc(x_248); lean_inc(x_247); -lean_inc(x_246); -lean_dec(x_242); -x_248 = l_Lean_IR_NormalizeIds_normIndex(x_238, x_2); +lean_dec(x_243); +x_249 = l_Lean_IR_NormalizeIds_normIndex(x_239, x_2); lean_dec(x_2); -lean_dec(x_238); -lean_ctor_set(x_1, 3, x_246); -lean_ctor_set(x_1, 1, x_248); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_1); -lean_ctor_set(x_249, 1, x_247); -return x_249; +lean_dec(x_239); +lean_ctor_set(x_1, 3, x_247); +lean_ctor_set(x_1, 1, x_249); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_1); +lean_ctor_set(x_250, 1, x_248); +return x_250; } } else { -lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; size_t x_254; size_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; -x_250 = lean_ctor_get(x_1, 0); -x_251 = lean_ctor_get(x_1, 1); -x_252 = lean_ctor_get(x_1, 2); -x_253 = lean_ctor_get(x_1, 3); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; size_t x_255; size_t x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_251 = lean_ctor_get(x_1, 0); +x_252 = lean_ctor_get(x_1, 1); +x_253 = lean_ctor_get(x_1, 2); +x_254 = lean_ctor_get(x_1, 3); +lean_inc(x_254); lean_inc(x_253); lean_inc(x_252); lean_inc(x_251); -lean_inc(x_250); lean_dec(x_1); -x_254 = lean_array_size(x_253); -x_255 = 0; +x_255 = lean_array_size(x_254); +x_256 = 0; lean_inc(x_2); -x_256 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__3(x_254, x_255, x_253, x_2, x_3); -x_257 = lean_ctor_get(x_256, 0); -lean_inc(x_257); -x_258 = lean_ctor_get(x_256, 1); +x_257 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_NormalizeIds_normFnBody_spec__3(x_255, x_256, x_254, x_2, x_3); +x_258 = lean_ctor_get(x_257, 0); lean_inc(x_258); -if (lean_is_exclusive(x_256)) { - lean_ctor_release(x_256, 0); - lean_ctor_release(x_256, 1); - x_259 = x_256; +x_259 = lean_ctor_get(x_257, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_260 = x_257; } else { - lean_dec_ref(x_256); - x_259 = lean_box(0); + lean_dec_ref(x_257); + x_260 = lean_box(0); } -x_260 = l_Lean_IR_NormalizeIds_normIndex(x_251, x_2); +x_261 = l_Lean_IR_NormalizeIds_normIndex(x_252, x_2); lean_dec(x_2); -lean_dec(x_251); -x_261 = lean_alloc_ctor(9, 4, 0); -lean_ctor_set(x_261, 0, x_250); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_261, 2, x_252); -lean_ctor_set(x_261, 3, x_257); -if (lean_is_scalar(x_259)) { - x_262 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_252); +x_262 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_262, 0, x_251); +lean_ctor_set(x_262, 1, x_261); +lean_ctor_set(x_262, 2, x_253); +lean_ctor_set(x_262, 3, x_258); +if (lean_is_scalar(x_260)) { + x_263 = lean_alloc_ctor(0, 2, 0); } else { - x_262 = x_259; + x_263 = x_260; } -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_258); -return x_262; +lean_ctor_set(x_263, 0, x_262); +lean_ctor_set(x_263, 1, x_259); +return x_263; } } case 10: { -uint8_t x_263; -x_263 = !lean_is_exclusive(x_1); -if (x_263 == 0) +uint8_t x_264; +x_264 = !lean_is_exclusive(x_1); +if (x_264 == 0) { -lean_object* x_264; lean_object* x_265; lean_object* x_266; -x_264 = lean_ctor_get(x_1, 0); -x_265 = l_Lean_IR_NormalizeIds_normArg(x_264, x_2); +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = lean_ctor_get(x_1, 0); +x_266 = l_Lean_IR_NormalizeIds_normArg(x_265, x_2); lean_dec(x_2); -lean_ctor_set(x_1, 0, x_265); -x_266 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_266, 0, x_1); -lean_ctor_set(x_266, 1, x_3); -return x_266; +lean_ctor_set(x_1, 0, x_266); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_1); +lean_ctor_set(x_267, 1, x_3); +return x_267; } else { -lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_267 = lean_ctor_get(x_1, 0); -lean_inc(x_267); +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +x_268 = lean_ctor_get(x_1, 0); +lean_inc(x_268); lean_dec(x_1); -x_268 = l_Lean_IR_NormalizeIds_normArg(x_267, x_2); +x_269 = l_Lean_IR_NormalizeIds_normArg(x_268, x_2); lean_dec(x_2); -x_269 = lean_alloc_ctor(10, 1, 0); -lean_ctor_set(x_269, 0, x_268); -x_270 = lean_alloc_ctor(0, 2, 0); +x_270 = lean_alloc_ctor(10, 1, 0); lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_3); -return x_270; +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_3); +return x_271; } } case 11: { -uint8_t x_271; -x_271 = !lean_is_exclusive(x_1); -if (x_271 == 0) +uint8_t x_272; +x_272 = !lean_is_exclusive(x_1); +if (x_272 == 0) { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_272 = lean_ctor_get(x_1, 0); -x_273 = lean_ctor_get(x_1, 1); -x_274 = l_Lean_IR_NormalizeIds_normIndex(x_272, x_2); -lean_dec(x_272); -x_275 = l_Lean_IR_NormalizeIds_normArgs(x_273, x_2); +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_273 = lean_ctor_get(x_1, 0); +x_274 = lean_ctor_get(x_1, 1); +x_275 = l_Lean_IR_NormalizeIds_normIndex(x_273, x_2); +lean_dec(x_273); +x_276 = l_Lean_IR_NormalizeIds_normArgs(x_274, x_2); lean_dec(x_2); -lean_ctor_set(x_1, 1, x_275); -lean_ctor_set(x_1, 0, x_274); -x_276 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_276, 0, x_1); -lean_ctor_set(x_276, 1, x_3); -return x_276; +lean_ctor_set(x_1, 1, x_276); +lean_ctor_set(x_1, 0, x_275); +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; } else { -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; -x_277 = lean_ctor_get(x_1, 0); -x_278 = lean_ctor_get(x_1, 1); +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_278 = lean_ctor_get(x_1, 0); +x_279 = lean_ctor_get(x_1, 1); +lean_inc(x_279); lean_inc(x_278); -lean_inc(x_277); lean_dec(x_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); +x_280 = l_Lean_IR_NormalizeIds_normIndex(x_278, x_2); +lean_dec(x_278); +x_281 = l_Lean_IR_NormalizeIds_normArgs(x_279, x_2); lean_dec(x_2); -x_281 = lean_alloc_ctor(11, 2, 0); -lean_ctor_set(x_281, 0, x_279); -lean_ctor_set(x_281, 1, x_280); -x_282 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_282, 0, x_281); -lean_ctor_set(x_282, 1, x_3); -return x_282; +x_282 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_282, 0, x_280); +lean_ctor_set(x_282, 1, x_281); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_3); +return x_283; } } default: { -lean_object* x_283; +lean_object* x_284; lean_dec(x_2); -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; +x_284 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_284, 0, x_1); +lean_ctor_set(x_284, 1, x_3); +return x_284; } } } @@ -4633,185 +4637,187 @@ return x_2; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; x_39 = lean_ctor_get(x_2, 0); -x_40 = lean_ctor_get(x_2, 1); -lean_inc(x_40); +x_40 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_41 = lean_ctor_get(x_2, 1); +lean_inc(x_41); 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_42 = lean_apply_1(x_1, x_41); +x_43 = lean_alloc_ctor(4, 2, 1); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*2, x_40); +return x_43; } } case 5: { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_2); -if (x_43 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_2); +if (x_44 == 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_45; lean_object* x_46; +x_45 = lean_ctor_get(x_2, 2); +x_46 = lean_apply_1(x_1, x_45); +lean_ctor_set(x_2, 2, x_46); 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_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_2, 0); +x_48 = lean_ctor_get(x_2, 1); +x_49 = lean_ctor_get(x_2, 2); +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_50 = lean_apply_1(x_1, x_49); +x_51 = lean_alloc_ctor(5, 3, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_48); +lean_ctor_set(x_51, 2, x_50); +return x_51; } } case 6: { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_2); -if (x_51 == 0) +uint8_t x_52; +x_52 = !lean_is_exclusive(x_2); +if (x_52 == 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_53; lean_object* x_54; +x_53 = lean_ctor_get(x_2, 1); +x_54 = l_Lean_IR_MapVars_mapArgs(x_1, x_53); +lean_ctor_set(x_2, 1, x_54); 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_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_2, 0); +x_56 = lean_ctor_get(x_2, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); 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_57 = l_Lean_IR_MapVars_mapArgs(x_1, x_56); +x_58 = lean_alloc_ctor(6, 2, 0); +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_57); +return x_58; } } case 7: { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_2); -if (x_58 == 0) +uint8_t x_59; +x_59 = !lean_is_exclusive(x_2); +if (x_59 == 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_60; lean_object* x_61; +x_60 = lean_ctor_get(x_2, 1); +x_61 = l_Lean_IR_MapVars_mapArgs(x_1, x_60); +lean_ctor_set(x_2, 1, x_61); 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_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_2, 0); +x_63 = lean_ctor_get(x_2, 1); +lean_inc(x_63); lean_inc(x_62); -lean_inc(x_61); 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_64 = l_Lean_IR_MapVars_mapArgs(x_1, x_63); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } case 8: { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_2); -if (x_65 == 0) +uint8_t x_66; +x_66 = !lean_is_exclusive(x_2); +if (x_66 == 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_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_2, 0); +x_68 = 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_69 = lean_apply_1(x_1, x_67); +x_70 = l_Lean_IR_MapVars_mapArgs(x_1, x_68); +lean_ctor_set(x_2, 1, x_70); +lean_ctor_set(x_2, 0, x_69); 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_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_2, 0); +x_72 = lean_ctor_get(x_2, 1); +lean_inc(x_72); lean_inc(x_71); -lean_inc(x_70); 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_73 = lean_apply_1(x_1, x_71); +x_74 = l_Lean_IR_MapVars_mapArgs(x_1, x_72); +x_75 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } case 9: { -uint8_t x_75; -x_75 = !lean_is_exclusive(x_2); -if (x_75 == 0) +uint8_t x_76; +x_76 = !lean_is_exclusive(x_2); +if (x_76 == 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_77; lean_object* x_78; +x_77 = lean_ctor_get(x_2, 1); +x_78 = lean_apply_1(x_1, x_77); +lean_ctor_set(x_2, 1, x_78); 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_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_79 = lean_ctor_get(x_2, 0); +x_80 = lean_ctor_get(x_2, 1); +lean_inc(x_80); lean_inc(x_79); -lean_inc(x_78); 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_81 = lean_apply_1(x_1, x_80); +x_82 = lean_alloc_ctor(9, 2, 0); +lean_ctor_set(x_82, 0, x_79); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } case 10: { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_2); -if (x_82 == 0) +uint8_t x_83; +x_83 = !lean_is_exclusive(x_2); +if (x_83 == 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_84; lean_object* x_85; +x_84 = lean_ctor_get(x_2, 0); +x_85 = lean_apply_1(x_1, x_84); +lean_ctor_set(x_2, 0, x_85); 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_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_2, 0); +lean_inc(x_86); 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_87 = lean_apply_1(x_1, x_86); +x_88 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_88, 0, x_87); +return x_88; } } case 11: @@ -4821,26 +4827,26 @@ return x_2; } default: { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_2); -if (x_88 == 0) +uint8_t x_89; +x_89 = !lean_is_exclusive(x_2); +if (x_89 == 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_90; lean_object* x_91; +x_90 = lean_ctor_get(x_2, 0); +x_91 = lean_apply_1(x_1, x_90); +lean_ctor_set(x_2, 0, x_91); 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_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_2, 0); +lean_inc(x_92); 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_93 = lean_apply_1(x_1, x_92); +x_94 = lean_alloc_ctor(12, 1, 0); +lean_ctor_set(x_94, 0, x_93); +return x_94; } } } @@ -5146,295 +5152,297 @@ 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; uint8_t 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); -lean_inc(x_64); +x_64 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); +x_65 = lean_ctor_get(x_2, 3); +lean_inc(x_65); 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_63); +x_68 = l_Lean_IR_MapVars_mapFnBody(x_1, x_65); +x_69 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_69, 0, x_66); +lean_ctor_set(x_69, 1, x_62); +lean_ctor_set(x_69, 2, x_67); +lean_ctor_set(x_69, 3, x_68); +lean_ctor_set_uint8(x_69, sizeof(void*)*4, x_64); +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, 3); +x_73 = lean_ctor_get(x_2, 5); 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, 5, x_76); +lean_ctor_set(x_2, 3, 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; +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); +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_83 = 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_84 = lean_apply_1(x_1, x_80); +x_85 = l_Lean_IR_MapVars_mapFnBody(x_1, x_82); +x_86 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_78); +lean_ctor_set(x_86, 2, x_79); +lean_ctor_set(x_86, 3, x_84); +lean_ctor_set(x_86, 4, x_81); +lean_ctor_set(x_86, 5, x_85); +return x_86; } } case 6: { -uint8_t x_86; -x_86 = !lean_is_exclusive(x_2); -if (x_86 == 0) +uint8_t x_87; +x_87 = !lean_is_exclusive(x_2); +if (x_87 == 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_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_2, 0); +x_89 = 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_90 = lean_apply_1(x_1, x_88); +x_91 = l_Lean_IR_MapVars_mapFnBody(x_1, x_89); +lean_ctor_set(x_2, 2, x_91); +lean_ctor_set(x_2, 0, x_90); 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_object* x_92; lean_object* x_93; uint8_t x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_92 = lean_ctor_get(x_2, 0); +x_93 = lean_ctor_get(x_2, 1); +x_94 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_95 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_96 = lean_ctor_get(x_2, 2); +lean_inc(x_96); +lean_inc(x_93); lean_inc(x_92); -lean_inc(x_91); 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_97 = lean_apply_1(x_1, x_92); +x_98 = l_Lean_IR_MapVars_mapFnBody(x_1, x_96); +x_99 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_93); +lean_ctor_set(x_99, 2, x_98); +lean_ctor_set_uint8(x_99, sizeof(void*)*3, x_94); +lean_ctor_set_uint8(x_99, sizeof(void*)*3 + 1, x_95); +return x_99; } } case 7: { -uint8_t x_99; -x_99 = !lean_is_exclusive(x_2); -if (x_99 == 0) +uint8_t x_100; +x_100 = !lean_is_exclusive(x_2); +if (x_100 == 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_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_101 = lean_ctor_get(x_2, 0); +x_102 = 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_103 = lean_apply_1(x_1, x_101); +x_104 = l_Lean_IR_MapVars_mapFnBody(x_1, x_102); +lean_ctor_set(x_2, 2, x_104); +lean_ctor_set(x_2, 0, x_103); 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_object* x_105; lean_object* x_106; uint8_t x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_2, 0); +x_106 = lean_ctor_get(x_2, 1); +x_107 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_108 = lean_ctor_get_uint8(x_2, sizeof(void*)*3 + 1); +x_109 = lean_ctor_get(x_2, 2); +lean_inc(x_109); +lean_inc(x_106); lean_inc(x_105); -lean_inc(x_104); 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_110 = lean_apply_1(x_1, x_105); +x_111 = l_Lean_IR_MapVars_mapFnBody(x_1, x_109); +x_112 = lean_alloc_ctor(7, 3, 2); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_106); +lean_ctor_set(x_112, 2, x_111); +lean_ctor_set_uint8(x_112, sizeof(void*)*3, x_107); +lean_ctor_set_uint8(x_112, sizeof(void*)*3 + 1, x_108); +return x_112; } } case 8: { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_2); -if (x_112 == 0) +uint8_t x_113; +x_113 = !lean_is_exclusive(x_2); +if (x_113 == 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_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_114 = lean_ctor_get(x_2, 0); +x_115 = 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_116 = lean_apply_1(x_1, x_114); +x_117 = l_Lean_IR_MapVars_mapFnBody(x_1, x_115); +lean_ctor_set(x_2, 1, x_117); +lean_ctor_set(x_2, 0, x_116); 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_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_118 = lean_ctor_get(x_2, 0); +x_119 = lean_ctor_get(x_2, 1); +lean_inc(x_119); lean_inc(x_118); -lean_inc(x_117); 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_120 = lean_apply_1(x_1, x_118); +x_121 = l_Lean_IR_MapVars_mapFnBody(x_1, x_119); +x_122 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; } } case 9: { -uint8_t x_122; -x_122 = !lean_is_exclusive(x_2); -if (x_122 == 0) +uint8_t x_123; +x_123 = !lean_is_exclusive(x_2); +if (x_123 == 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_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_124 = lean_ctor_get(x_2, 1); +x_125 = 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_126 = lean_apply_1(x_1, x_124); +x_127 = lean_array_size(x_125); +x_128 = 0; +x_129 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_127, x_128, x_125); +lean_ctor_set(x_2, 3, x_129); +lean_ctor_set(x_2, 1, x_126); 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_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; size_t x_135; size_t x_136; lean_object* x_137; lean_object* x_138; +x_130 = lean_ctor_get(x_2, 0); +x_131 = lean_ctor_get(x_2, 1); +x_132 = lean_ctor_get(x_2, 2); +x_133 = lean_ctor_get(x_2, 3); +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_134 = lean_apply_1(x_1, x_131); +x_135 = lean_array_size(x_133); +x_136 = 0; +x_137 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_MapVars_mapFnBody_spec__0(x_1, x_135, x_136, x_133); +x_138 = lean_alloc_ctor(9, 4, 0); +lean_ctor_set(x_138, 0, x_130); +lean_ctor_set(x_138, 1, x_134); +lean_ctor_set(x_138, 2, x_132); +lean_ctor_set(x_138, 3, x_137); +return x_138; } } 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; -x_140 = lean_ctor_get(x_2, 0); -lean_dec(x_140); -x_141 = !lean_is_exclusive(x_138); -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_139; +x_139 = lean_ctor_get(x_2, 0); +lean_inc(x_139); +if (lean_obj_tag(x_139) == 0) +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_2); +if (x_140 == 0) +{ +lean_object* x_141; uint8_t x_142; +x_141 = lean_ctor_get(x_2, 0); +lean_dec(x_141); +x_142 = !lean_is_exclusive(x_139); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; +x_143 = lean_ctor_get(x_139, 0); +x_144 = lean_apply_1(x_1, x_143); +lean_ctor_set(x_139, 0, x_144); 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_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_139, 0); +lean_inc(x_145); +lean_dec(x_139); +x_146 = lean_apply_1(x_1, x_145); +x_147 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_2, 0, x_147); 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_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; 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_148 = lean_ctor_get(x_139, 0); +lean_inc(x_148); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + x_149 = x_139; } else { - lean_dec_ref(x_138); - x_148 = lean_box(0); + lean_dec_ref(x_139); + x_149 = 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_150 = lean_apply_1(x_1, x_148); +if (lean_is_scalar(x_149)) { + x_151 = lean_alloc_ctor(0, 1, 0); } else { - x_150 = x_148; + x_151 = x_149; } -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; +x_152 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_152, 0, x_151); +return x_152; } } else @@ -5445,29 +5453,29 @@ return x_2; } case 11: { -uint8_t x_152; -x_152 = !lean_is_exclusive(x_2); -if (x_152 == 0) +uint8_t x_153; +x_153 = !lean_is_exclusive(x_2); +if (x_153 == 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_154; lean_object* x_155; +x_154 = lean_ctor_get(x_2, 1); +x_155 = l_Lean_IR_MapVars_mapArgs(x_1, x_154); +lean_ctor_set(x_2, 1, x_155); 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_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_156 = lean_ctor_get(x_2, 0); +x_157 = lean_ctor_get(x_2, 1); +lean_inc(x_157); lean_inc(x_156); -lean_inc(x_155); 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_158 = l_Lean_IR_MapVars_mapArgs(x_1, x_157); +x_159 = lean_alloc_ctor(11, 2, 0); +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_158); +return x_159; } } default: diff --git a/stage0/stdlib/Lean/Compiler/IR/RC.c b/stage0/stdlib/Lean/Compiler/IR/RC.c index 402e8190e928..6d8aebb27418 100644 --- a/stage0/stdlib/Lean/Compiler/IR/RC.c +++ b/stage0/stdlib/Lean/Compiler/IR/RC.c @@ -9316,15 +9316,15 @@ uint8_t x_33; x_33 = lean_ctor_get_uint8(x_22, 1); if (x_33 == 0) { -x_15 = x_10; -x_16 = x_22; +x_15 = x_22; +x_16 = x_10; x_17 = x_27; goto block_21; } else { -x_15 = x_10; -x_16 = x_22; +x_15 = x_22; +x_16 = x_10; x_17 = x_32; goto block_21; } @@ -9348,12 +9348,12 @@ goto _start; block_21: { uint8_t x_18; lean_object* x_19; -x_18 = lean_ctor_get_uint8(x_16, 2); -lean_dec_ref(x_16); +x_18 = lean_ctor_get_uint8(x_15, 2); +lean_dec_ref(x_15); lean_inc(x_14); x_19 = lean_alloc_ctor(7, 3, 2); lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_15); +lean_ctor_set(x_19, 1, x_16); lean_ctor_set(x_19, 2, x_7); lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_17); lean_ctor_set_uint8(x_19, sizeof(void*)*3 + 1, x_18); @@ -9450,15 +9450,15 @@ uint8_t x_33; x_33 = lean_ctor_get_uint8(x_22, 1); if (x_33 == 0) { -x_15 = x_22; -x_16 = x_10; +x_15 = x_10; +x_16 = x_22; x_17 = x_27; goto block_21; } else { -x_15 = x_22; -x_16 = x_10; +x_15 = x_10; +x_16 = x_22; x_17 = x_32; goto block_21; } @@ -9484,12 +9484,12 @@ return x_35; block_21: { uint8_t x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get_uint8(x_15, 2); -lean_dec_ref(x_15); +x_18 = lean_ctor_get_uint8(x_16, 2); +lean_dec_ref(x_16); lean_inc(x_14); x_19 = lean_alloc_ctor(7, 3, 2); lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_16); +lean_ctor_set(x_19, 1, x_15); lean_ctor_set(x_19, 2, x_7); lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_17); lean_ctor_set_uint8(x_19, sizeof(void*)*3 + 1, x_18); @@ -10202,11 +10202,11 @@ return x_3; block_10: { uint8_t x_8; lean_object* x_9; -x_8 = lean_ctor_get_uint8(x_5, 2); -lean_dec_ref(x_5); +x_8 = lean_ctor_get_uint8(x_6, 2); +lean_dec_ref(x_6); x_9 = lean_alloc_ctor(7, 3, 2); lean_ctor_set(x_9, 0, x_2); -lean_ctor_set(x_9, 1, x_6); +lean_ctor_set(x_9, 1, x_5); lean_ctor_set(x_9, 2, x_3); lean_ctor_set_uint8(x_9, sizeof(void*)*3, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 1, x_8); @@ -10232,15 +10232,15 @@ x_15 = lean_ctor_get_uint8(x_11, 1); x_16 = lean_unsigned_to_nat(1u); if (x_15 == 0) { -x_5 = x_11; -x_6 = x_16; +x_5 = x_16; +x_6 = x_11; x_7 = x_12; goto block_10; } else { -x_5 = x_11; -x_6 = x_16; +x_5 = x_16; +x_6 = x_11; x_7 = x_14; goto block_10; } @@ -11219,9 +11219,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 { @@ -11240,7 +11240,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; } @@ -11266,8 +11266,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; } @@ -11275,8 +11275,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; } @@ -12793,139 +12793,141 @@ 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; uint8_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; x_93 = lean_ctor_get(x_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); -lean_inc(x_96); +x_96 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_97 = lean_ctor_get(x_1, 3); +lean_inc(x_97); 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, 4, 1); +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_99); +lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_96); +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, 5); 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) -{ -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); +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_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, 5, 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, 5, 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; +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); +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); +x_123 = l_Lean_IR_ExplicitRC_visitFnBody(x_122, x_2); +x_124 = lean_ctor_get(x_123, 0); 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_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_126 = x_123; } else { - lean_dec_ref(x_122); - x_125 = lean_box(0); + lean_dec_ref(x_123); + x_126 = 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_127 = 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_125); 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); +x_128 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_128, 0, x_117); +lean_ctor_set(x_128, 1, x_118); +lean_ctor_set(x_128, 2, x_119); +lean_ctor_set(x_128, 3, x_120); +lean_ctor_set(x_128, 4, x_121); +lean_ctor_set(x_128, 5, x_124); +if (lean_is_scalar(x_126)) { + x_129 = lean_alloc_ctor(0, 2, 0); } else { - x_128 = x_125; + x_129 = x_126; } -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -return x_128; +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +return x_129; } } case 6: @@ -12942,212 +12944,212 @@ 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_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; size_t x_135; size_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +x_130 = lean_ctor_get(x_1, 0); lean_inc(x_130); -x_131 = lean_ctor_get(x_1, 2); +x_131 = lean_ctor_get(x_1, 1); lean_inc(x_131); -x_132 = lean_ctor_get(x_1, 3); -lean_inc_ref(x_132); +x_132 = lean_ctor_get(x_1, 2); +lean_inc(x_132); +x_133 = lean_ctor_get(x_1, 3); +lean_inc_ref(x_133); 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_134 = x_1; } else { lean_dec_ref(x_1); - x_133 = lean_box(0); + x_134 = lean_box(0); } -x_134 = lean_array_size(x_132); -x_135 = 0; -lean_inc(x_130); +x_135 = lean_array_size(x_133); +x_136 = 0; +lean_inc(x_131); lean_inc_ref(x_2); -x_136 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__5(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_137 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__5(x_2, x_131, x_135, x_136, x_133); +x_145 = l_Lean_IR_ExplicitRC_instInhabitedLiveVars_default___closed__0; +x_146 = lean_unsigned_to_nat(0u); +x_147 = lean_array_get_size(x_137); +x_148 = lean_nat_dec_lt(x_146, x_147); +if (x_148 == 0) { -lean_dec(x_146); -x_137 = x_144; -goto block_143; +lean_dec(x_147); +x_138 = x_145; +goto block_144; } else { -uint8_t x_148; -x_148 = lean_nat_dec_le(x_146, x_146); -if (x_148 == 0) +uint8_t x_149; +x_149 = lean_nat_dec_le(x_147, x_147); +if (x_149 == 0) { -lean_dec(x_146); -x_137 = x_144; -goto block_143; +lean_dec(x_147); +x_138 = x_145; +goto block_144; } else { -size_t x_149; lean_object* x_150; -x_149 = lean_usize_of_nat(x_146); -lean_dec(x_146); -x_150 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__7(x_136, x_135, x_149, x_144); -x_137 = x_150; -goto block_143; +size_t x_150; lean_object* x_151; +x_150 = lean_usize_of_nat(x_147); +lean_dec(x_147); +x_151 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_IR_ExplicitRC_visitFnBody_spec__7(x_137, x_136, x_150, x_145); +x_138 = x_151; +goto block_144; } } -block_143: +block_144: { -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__6(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); -} else { - x_141 = x_133; -} -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_object* x_139; size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +lean_inc(x_131); +x_139 = 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_131, x_138); +x_140 = lean_array_size(x_137); +lean_inc(x_131); +x_141 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_IR_ExplicitRC_visitFnBody_spec__6(x_2, x_131, x_139, x_140, x_136, x_137); +if (lean_is_scalar(x_134)) { + x_142 = lean_alloc_ctor(9, 4, 0); +} else { + x_142 = x_134; +} +lean_ctor_set(x_142, 0, x_130); +lean_ctor_set(x_142, 1, x_131); +lean_ctor_set(x_142, 2, x_132); +lean_ctor_set(x_142, 3, x_141); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_139); +return x_143; } } case 10: { -lean_object* x_151; lean_object* x_152; -x_151 = lean_ctor_get(x_1, 0); +lean_object* x_152; lean_object* x_153; +x_152 = lean_ctor_get(x_1, 0); lean_inc_ref(x_2); -x_152 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); -if (lean_obj_tag(x_151) == 0) -{ -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; -x_153 = lean_ctor_get(x_151, 0); -x_154 = lean_ctor_get(x_2, 4); -x_155 = l_Lean_IR_ExplicitRC_instInhabitedVarInfo_default; -x_156 = l_Std_DTreeMap_Internal_Impl_Const_get_x21___at___00Lean_IR_ExplicitRC_getVarInfo_spec__0___redArg(x_154, x_153, x_155); -x_157 = lean_ctor_get_uint8(x_156, 0); -lean_dec_ref(x_156); -lean_inc(x_153); -x_158 = 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_153, x_152); -if (x_157 == 0) +x_153 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); +if (lean_obj_tag(x_152) == 0) +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; +x_154 = lean_ctor_get(x_152, 0); +x_155 = lean_ctor_get(x_2, 4); +x_156 = l_Lean_IR_ExplicitRC_instInhabitedVarInfo_default; +x_157 = l_Std_DTreeMap_Internal_Impl_Const_get_x21___at___00Lean_IR_ExplicitRC_getVarInfo_spec__0___redArg(x_155, x_154, x_156); +x_158 = lean_ctor_get_uint8(x_157, 0); +lean_dec_ref(x_157); +lean_inc(x_154); +x_159 = 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_154, x_153); +if (x_158 == 0) { -lean_object* x_166; +lean_object* x_167; lean_dec_ref(x_2); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_1); -lean_ctor_set(x_166, 1, x_158); -return x_166; +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_1); +lean_ctor_set(x_167, 1, x_159); +return x_167; } else { -lean_object* x_167; uint8_t x_168; -x_167 = lean_ctor_get(x_158, 1); -lean_inc(x_167); -x_168 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__6___redArg(x_153, x_167); -lean_dec(x_167); -if (x_168 == 0) +lean_object* x_168; uint8_t x_169; +x_168 = lean_ctor_get(x_159, 1); +lean_inc(x_168); +x_169 = l_Std_DTreeMap_Internal_Impl_contains___at___00__private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_CollectDerivedValInfo_visitParam_spec__6___redArg(x_154, x_168); +lean_dec(x_168); +if (x_169 == 0) { -lean_object* x_169; +lean_object* x_170; lean_dec_ref(x_2); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_1); -lean_ctor_set(x_169, 1, x_158); -return x_169; +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_1); +lean_ctor_set(x_170, 1, x_159); +return x_170; } else { -lean_object* x_170; uint8_t x_171; lean_object* x_172; -lean_inc(x_153); -x_170 = l_Lean_IR_ExplicitRC_getVarInfo(x_2, x_153); +lean_object* x_171; uint8_t x_172; lean_object* x_173; +lean_inc(x_154); +x_171 = l_Lean_IR_ExplicitRC_getVarInfo(x_2, x_154); lean_dec_ref(x_2); -x_171 = lean_ctor_get_uint8(x_170, 1); -x_172 = lean_unsigned_to_nat(1u); -if (x_171 == 0) +x_172 = lean_ctor_get_uint8(x_171, 1); +x_173 = lean_unsigned_to_nat(1u); +if (x_172 == 0) { -x_159 = x_170; -x_160 = x_172; -x_161 = x_168; -goto block_165; +x_160 = x_171; +x_161 = x_173; +x_162 = x_169; +goto block_166; } else { -uint8_t x_173; -x_173 = 0; -x_159 = x_170; -x_160 = x_172; +uint8_t x_174; +x_174 = 0; +x_160 = x_171; x_161 = x_173; -goto block_165; +x_162 = x_174; +goto block_166; } } } -block_165: +block_166: { -uint8_t x_162; lean_object* x_163; lean_object* x_164; -x_162 = lean_ctor_get_uint8(x_159, 2); -lean_dec_ref(x_159); -x_163 = lean_alloc_ctor(6, 3, 2); -lean_ctor_set(x_163, 0, x_153); -lean_ctor_set(x_163, 1, x_160); -lean_ctor_set(x_163, 2, x_1); -lean_ctor_set_uint8(x_163, sizeof(void*)*3, x_161); -lean_ctor_set_uint8(x_163, sizeof(void*)*3 + 1, x_162); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_158); -return x_164; +uint8_t x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get_uint8(x_160, 2); +lean_dec_ref(x_160); +x_164 = lean_alloc_ctor(6, 3, 2); +lean_ctor_set(x_164, 0, x_154); +lean_ctor_set(x_164, 1, x_161); +lean_ctor_set(x_164, 2, x_1); +lean_ctor_set_uint8(x_164, sizeof(void*)*3, x_162); +lean_ctor_set_uint8(x_164, sizeof(void*)*3 + 1, x_163); +x_165 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_159); +return x_165; } } else { -lean_object* x_174; +lean_object* x_175; 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_152); -return x_174; +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_1); +lean_ctor_set(x_175, 1, x_153); +return x_175; } } 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_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; +x_176 = lean_ctor_get(x_1, 0); +x_177 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_177); +x_178 = l_Lean_IR_ExplicitRC_getJPLiveVars(x_2, x_176); +x_179 = l_Lean_IR_ExplicitRC_getJPParams(x_2, x_176); +x_180 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_addIncBefore(x_2, x_177, x_179, x_1, x_178); +x_181 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_useArgs(x_2, x_177, x_178); +lean_dec_ref(x_177); 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; +x_182 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; } 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, 1, x_182); -return x_183; +lean_object* x_183; lean_object* x_184; +x_183 = l___private_Lean_Compiler_IR_RC_0__Lean_IR_ExplicitRC_mkRetLiveVars(x_2); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_1); +lean_ctor_set(x_184, 1, x_183); +return x_184; } default: { -lean_object* x_184; lean_object* x_185; +lean_object* x_185; lean_object* x_186; 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__3(x_184); -return x_185; +x_185 = l_Lean_IR_ExplicitRC_visitFnBody___closed__3; +x_186 = l_panic___at___00Lean_IR_ExplicitRC_visitFnBody_spec__3(x_185); +return x_186; } } block_5: diff --git a/stage0/stdlib/Lean/Compiler/IR/ToIR.c b/stage0/stdlib/Lean/Compiler/IR/ToIR.c index 3fa57bc8dd9e..ed93b2e26c7f 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ToIR.c +++ b/stage0/stdlib/Lean/Compiler/IR/ToIR.c @@ -217,6 +217,7 @@ lean_object* l_Lean_IR_IRType_boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ToIR_M_run___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ToIR_TranslatedProj_erased_elim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___00__private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerAlt_loop_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_IRType_ptrSizedTypeForSign(uint8_t); static lean_object* l_Lean_IR_ToIR_lowerCode___closed__6; lean_object* l_Lean_MessageData_tagWithErrorName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_mkVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2467,93 +2468,75 @@ return x_13; } case 2: { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_3); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -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_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_3, 0); +lean_inc(x_14); +x_15 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +lean_dec_ref(x_3); +x_16 = lean_alloc_ctor(4, 2, 1); +lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, 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; -} -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_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; -} +lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_15); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = l_Lean_IR_IRType_ptrSizedTypeForSign(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } case 3: { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_3); -if (x_24 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_3); +if (x_20 == 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_object* x_21; 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_3, 2); +x_22 = lean_ctor_get(x_3, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_2, 2); +x_24 = lean_ctor_get(x_2, 3); +x_25 = lean_nat_add(x_23, x_24); 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; +lean_ctor_set(x_3, 0, x_25); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_3); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_21); +return x_27; } 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_object* x_28; lean_object* x_29; 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_3, 1); +x_29 = lean_ctor_get(x_3, 2); +lean_inc(x_29); +lean_inc(x_28); 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; +x_30 = lean_ctor_get(x_2, 2); +x_31 = lean_ctor_get(x_2, 3); +x_32 = lean_nat_add(x_30, x_31); +x_33 = lean_alloc_ctor(5, 3, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_28); +lean_ctor_set(x_33, 2, x_1); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, 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_29); +return x_35; } } default: { -lean_object* x_40; +lean_object* x_36; lean_dec(x_1); -x_40 = l_Lean_IR_ToIR_lowerProj___closed__1; -return x_40; +x_36 = l_Lean_IR_ToIR_lowerProj___closed__1; +return x_36; } } } @@ -8007,137 +7990,140 @@ goto _start; } case 2: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_6, x_22); +x_22 = lean_ctor_get_uint8(x_17, sizeof(void*)*1); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_6, x_23); lean_dec(x_6); lean_inc(x_5); -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) +x_25 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_24, x_7, x_8, x_9); +if (lean_obj_tag(x_25) == 0) { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); 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); -return x_24; +x_28 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_28, 0, x_5); +lean_ctor_set(x_28, 1, x_21); +lean_ctor_set(x_28, 2, x_15); +lean_ctor_set(x_28, 3, x_27); +lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_22); +lean_ctor_set(x_25, 0, x_28); +return x_25; } 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_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_25, 0); +lean_inc(x_29); +lean_dec(x_25); 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_30 = lean_alloc_ctor(4, 4, 1); +lean_ctor_set(x_30, 0, x_5); +lean_ctor_set(x_30, 1, x_21); +lean_ctor_set(x_30, 2, x_15); +lean_ctor_set(x_30, 3, x_29); +lean_ctor_set_uint8(x_30, sizeof(void*)*4, x_22); +x_31 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; } } else { lean_dec(x_5); -return x_24; +return x_25; } } 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_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_32 = lean_ctor_get(x_17, 1); +x_33 = lean_ctor_get(x_17, 2); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_6, x_34); 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) +x_36 = l___private_Lean_Compiler_IR_ToIR_0__Lean_IR_ToIR_lowerLet_lowerNonObjectFields_loop(x_1, x_2, x_3, x_4, x_5, x_35, x_7, x_8, x_9); +if (lean_obj_tag(x_36) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = 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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_2, 2); +x_40 = lean_ctor_get(x_2, 3); +x_41 = lean_nat_add(x_39, x_40); +lean_inc(x_33); 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_32); +x_42 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_42, 0, x_5); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_32); +lean_ctor_set(x_42, 3, x_15); +lean_ctor_set(x_42, 4, x_33); +lean_ctor_set(x_42, 5, x_38); +lean_ctor_set(x_36, 0, x_42); +return x_36; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_42 = lean_ctor_get(x_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_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_43 = lean_ctor_get(x_36, 0); +lean_inc(x_43); +lean_dec(x_36); +x_44 = lean_ctor_get(x_2, 2); +x_45 = lean_ctor_get(x_2, 3); +x_46 = lean_nat_add(x_44, x_45); +lean_inc(x_33); 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_32); +x_47 = lean_alloc_ctor(5, 6, 0); +lean_ctor_set(x_47, 0, x_5); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_32); +lean_ctor_set(x_47, 3, x_15); +lean_ctor_set(x_47, 4, x_33); +lean_ctor_set(x_47, 5, x_43); +x_48 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_48, 0, x_47); +return x_48; } } else { lean_dec(x_5); -return x_35; +return x_36; } } 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_49; lean_object* x_50; +x_49 = lean_unsigned_to_nat(1u); +x_50 = lean_nat_add(x_6, x_49); lean_dec(x_6); -x_6 = x_49; +x_6 = x_50; 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_52; lean_object* x_53; +x_52 = lean_unsigned_to_nat(1u); +x_53 = lean_nat_add(x_6, x_52); lean_dec(x_6); -x_6 = x_52; +x_6 = x_53; goto _start; } } diff --git a/stage0/stdlib/Lean/Compiler/IR/ToIRType.c b/stage0/stdlib/Lean/Compiler/IR/ToIRType.c index d004e3b830b4..15a7fbfa62d2 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ToIRType.c +++ b/stage0/stdlib/Lean/Compiler/IR/ToIRType.c @@ -19,7 +19,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___00__private_Lean_Com LEAN_EXPORT lean_object* l_Lean_IR_nameToIRType___boxed(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(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_IR_initFn_00___x40_Lean_Compiler_IR_ToIRType_3306651989____hygCtx___hyg_2__spec__0___redArg___closed__0; -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTR___redArg(lean_object*); static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__10; LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object*, lean_object*); @@ -35,7 +35,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCto static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__4; lean_object* l_Lean_Expr_eqv___boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__2; -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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*); +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static 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__9_spec__9___redArg___closed__0; lean_object* l_Lean_instHashableLevelMVarId_hash___boxed(lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); @@ -49,6 +49,7 @@ static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRTyp size_t lean_uint64_to_usize(uint64_t); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__22; static lean_object* l_Lean_IR_toIRType___closed__0; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31; 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__0_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instHashableInfoCacheKey___private__1___boxed(lean_object*); extern lean_object* l_Lean_Meta_instInhabitedConfigWithKey___private__1; @@ -125,6 +126,8 @@ lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp 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*); 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__9___redArg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__17; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30; 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; @@ -136,6 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_scalar_elim___redArg(lean_objec 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*); +static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__15; static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__11; static lean_object* l_Lean_IR_hasTrivialStructure_x3f___lam__0___closed__1; 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__8___redArg(lean_object*, lean_object*); @@ -146,6 +150,7 @@ static lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00Lean_I 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*); +static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__14; 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__0_spec__2___redArg(size_t, 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*); 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**); @@ -199,6 +204,7 @@ lean_object* l_Lean_Compiler_LCNF_getOtherDeclBaseType(lean_object*, lean_object static lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__3; 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___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28; lean_object* l___private_Init_Data_List_Impl_0__List_takeTR_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_CtorFieldInfo_format___closed__10; static 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__9_spec__9___redArg___closed__2; @@ -208,6 +214,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_CacheExtension_register___at___00L uint8_t l_Lean_Expr_isErased(lean_object*); lean_object* lean_usize_to_nat(size_t); static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__11; +static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__32; 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*); static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__15; 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__0_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -320,7 +327,7 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__9; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object*, 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); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_object_elim___redArg(lean_object*, lean_object*); @@ -332,7 +339,7 @@ static lean_object* l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRTyp 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*); -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorFieldInfo_ctorElim(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___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*); @@ -3241,7 +3248,7 @@ static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___a lean_object* 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___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__2; x_2 = lean_unsigned_to_nat(62u); -x_3 = lean_unsigned_to_nat(81u); +x_3 = lean_unsigned_to_nat(86u); x_4 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__1; x_5 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -4226,7 +4233,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameT _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float", 5, 5); +x_1 = lean_mk_string_unchecked("Int8", 4, 4); return x_1; } } @@ -4234,7 +4241,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameT _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float32", 7, 7); +x_1 = lean_mk_string_unchecked("Int16", 5, 5); return x_1; } } @@ -4242,7 +4249,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameT _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("lcErased", 8, 8); +x_1 = lean_mk_string_unchecked("Int32", 5, 5); return x_1; } } @@ -4250,7 +4257,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameT _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Int", 3, 3); +x_1 = lean_mk_string_unchecked("Int64", 5, 5); return x_1; } } @@ -4258,6 +4265,46 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameT _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("ISize", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Float", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Float32", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lcErased", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31() { +_start: +{ +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__32() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("lcVoid", 6, 6); return x_1; } @@ -4342,6 +4389,31 @@ x_82 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache__ x_83 = lean_string_dec_eq(x_63, x_82); if (x_83 == 0) { +lean_object* x_84; uint8_t x_85; +x_84 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28; +x_85 = lean_string_dec_eq(x_63, x_84); +if (x_85 == 0) +{ +lean_object* x_86; uint8_t x_87; +x_86 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29; +x_87 = lean_string_dec_eq(x_63, x_86); +if (x_87 == 0) +{ +lean_object* x_88; uint8_t x_89; +x_88 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30; +x_89 = lean_string_dec_eq(x_63, x_88); +if (x_89 == 0) +{ +lean_object* x_90; uint8_t x_91; +x_90 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31; +x_91 = lean_string_dec_eq(x_63, x_90); +if (x_91 == 0) +{ +lean_object* x_92; uint8_t x_93; +x_92 = l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__32; +x_93 = lean_string_dec_eq(x_63, x_92); +if (x_93 == 0) +{ x_28 = x_2; x_29 = x_3; x_30 = lean_box(0); @@ -4349,122 +4421,182 @@ goto block_61; } else { -lean_object* x_84; lean_object* x_85; +lean_object* x_94; lean_object* x_95; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_84 = lean_box(13); -x_85 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_85, 0, x_84); -return x_85; +x_94 = lean_box(13); +x_95 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_95, 0, x_94); +return x_95; } } else { -lean_object* x_86; lean_object* x_87; +lean_object* x_96; lean_object* x_97; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_86 = lean_box(8); -x_87 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_87, 0, x_86); -return x_87; +x_96 = lean_box(8); +x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; } } else { -lean_object* x_88; lean_object* x_89; +lean_object* x_98; lean_object* x_99; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_88 = lean_box(6); -x_89 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; +x_98 = lean_box(6); +x_99 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_99, 0, x_98); +return x_99; } } else { -lean_object* x_90; lean_object* x_91; +lean_object* x_100; lean_object* x_101; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_90 = lean_box(9); -x_91 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_91, 0, x_90); -return x_91; +x_100 = lean_box(9); +x_101 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_101, 0, x_100); +return x_101; +} +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_102 = lean_box(0); +x_103 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_103, 0, x_102); +return x_103; } } else { -lean_object* x_92; lean_object* x_93; +lean_object* x_104; lean_object* x_105; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_92 = lean_box(0); -x_93 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_93, 0, x_92); -return x_93; +x_104 = lean_box(18); +x_105 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_105, 0, x_104); +return x_105; } } else { -lean_object* x_94; lean_object* x_95; +lean_object* x_106; lean_object* x_107; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_94 = lean_box(5); -x_95 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_95, 0, x_94); -return x_95; +x_106 = lean_box(17); +x_107 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_107, 0, x_106); +return x_107; } } else { -lean_object* x_96; lean_object* x_97; +lean_object* x_108; lean_object* x_109; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_96 = lean_box(4); -x_97 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_97, 0, x_96); -return x_97; +x_108 = lean_box(16); +x_109 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_109, 0, x_108); +return x_109; } } else { -lean_object* x_98; lean_object* x_99; +lean_object* x_110; lean_object* x_111; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_98 = lean_box(3); -x_99 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_99, 0, x_98); -return x_99; +x_110 = lean_box(15); +x_111 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_111, 0, x_110); +return x_111; } } else { -lean_object* x_100; lean_object* x_101; +lean_object* x_112; lean_object* x_113; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_100 = lean_box(2); -x_101 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_101, 0, x_100); -return x_101; +x_112 = lean_box(14); +x_113 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_113, 0, x_112); +return x_113; } } else { -lean_object* x_102; lean_object* x_103; +lean_object* x_114; lean_object* x_115; lean_dec(x_3); lean_dec_ref(x_2); lean_dec_ref(x_1); -x_102 = lean_box(1); -x_103 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_103, 0, x_102); -return x_103; +x_114 = lean_box(5); +x_115 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_115, 0, x_114); +return x_115; +} +} +else +{ +lean_object* x_116; lean_object* x_117; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_116 = lean_box(4); +x_117 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_117, 0, x_116); +return x_117; +} +} +else +{ +lean_object* x_118; lean_object* x_119; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_118 = lean_box(3); +x_119 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_119, 0, x_118); +return x_119; +} +} +else +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_120 = lean_box(2); +x_121 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_121, 0, x_120); +return x_121; +} +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_122 = lean_box(1); +x_123 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_123, 0, x_122); +return x_123; } } else @@ -5575,7 +5707,7 @@ static lean_object* _init_l_Lean_IR_toIRType___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_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__2; x_2 = lean_unsigned_to_nat(41u); -x_3 = lean_unsigned_to_nat(110u); +x_3 = lean_unsigned_to_nat(115u); x_4 = l_Lean_IR_toIRType___closed__2; x_5 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -5588,7 +5720,7 @@ static lean_object* _init_l_Lean_IR_toIRType___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_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__2; x_2 = lean_unsigned_to_nat(9u); -x_3 = lean_unsigned_to_nat(122u); +x_3 = lean_unsigned_to_nat(127u); x_4 = l_Lean_IR_toIRType___closed__2; x_5 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -5774,25 +5906,27 @@ return x_5; } case 2: { -lean_object* x_6; lean_object* x_7; +lean_object* x_6; uint8_t 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_uint8(x_1, sizeof(void*)*1); lean_dec_ref(x_1); -x_7 = lean_apply_1(x_2, x_6); -return x_7; +x_8 = lean_box(x_7); +x_9 = lean_apply_2(x_2, x_6, x_8); +return x_9; } 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_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 0); lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 2); +lean_inc(x_12); lean_dec_ref(x_1); -x_11 = lean_apply_3(x_2, x_8, x_9, x_10); -return x_11; +x_13 = lean_apply_3(x_2, x_10, x_11, x_12); +return x_13; } default: { @@ -5991,7 +6125,7 @@ static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__8() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("scalar#", 7, 7); +x_1 = lean_mk_string_unchecked("isize@", 6, 6); return x_1; } } @@ -6009,7 +6143,7 @@ static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("@", 1, 1); +x_1 = lean_mk_string_unchecked("scalar#", 7, 7); return x_1; } } @@ -6027,7 +6161,7 @@ static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("void", 4, 4); +x_1 = lean_mk_string_unchecked("@", 1, 1); return x_1; } } @@ -6041,6 +6175,24 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("void", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_CtorFieldInfo_format___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_IR_CtorFieldInfo_format___closed__14; +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_CtorFieldInfo_format(lean_object* x_1) { _start: { @@ -6106,78 +6258,80 @@ return x_22; case 2: { uint8_t x_23; -x_23 = !lean_is_exclusive(x_1); +x_23 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); if (x_23 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_1, 0); +lean_inc(x_24); +lean_dec_ref(x_1); 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; +x_27 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_27); +return x_28; } 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_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; +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); +lean_dec_ref(x_1); +x_30 = l_Lean_IR_CtorFieldInfo_format___closed__9; +x_31 = l_Nat_reprFast(x_29); +x_32 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_33, 0, x_30); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } 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_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_34 = lean_ctor_get(x_1, 0); lean_inc(x_34); -x_35 = lean_ctor_get(x_1, 2); +x_35 = lean_ctor_get(x_1, 1); lean_inc(x_35); +x_36 = lean_ctor_get(x_1, 2); +lean_inc(x_36); 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; +x_37 = l_Lean_IR_CtorFieldInfo_format___closed__11; +x_38 = l_Nat_reprFast(x_34); +x_39 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_IR_CtorFieldInfo_format___closed__13; +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Nat_reprFast(x_35); +x_44 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_IR_CtorFieldInfo_format___closed__5; +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_IR_instToFormatIRType___private__1(x_36); +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } default: { -lean_object* x_49; -x_49 = l_Lean_IR_CtorFieldInfo_format___closed__13; -return x_49; +lean_object* x_50; +x_50 = l_Lean_IR_CtorFieldInfo_format___closed__15; +return x_50; } } } @@ -6452,2086 +6606,2497 @@ static lean_object* _init_l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at_ lean_object* 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___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__2; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(208u); +x_3 = lean_unsigned_to_nat(226u); x_4 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; x_5 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_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_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___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_10; -x_10 = lean_ctor_get(x_3, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) -{ lean_object* x_11; -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); -lean_dec_ref(x_3); +x_11 = lean_ctor_get(x_4, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +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(x_2); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_4); -return x_11; +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_5); +return x_12; } else { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_3); -if (x_12 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_3, 1); -x_14 = lean_ctor_get(x_3, 0); -lean_dec(x_14); -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_4, 1); +x_15 = lean_ctor_get(x_4, 0); +lean_dec(x_15); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_nat_dec_lt(x_16, x_13); -if (x_17 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_nat_dec_lt(x_17, x_14); +if (x_18 == 0) { -lean_object* x_18; -lean_free_object(x_10); -lean_dec(x_16); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_19; +lean_free_object(x_11); +lean_dec(x_17); +lean_free_object(x_4); +lean_dec(x_14); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_18 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_18, 0, x_4); -return x_18; +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_5); +return x_19; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_1, 0); -x_20 = lean_array_fget_borrowed(x_19, x_16); -x_21 = l_Lean_Expr_fvarId_x21(x_20); -lean_inc_ref(x_5); -x_22 = l_Lean_FVarId_getType___redArg(x_21, x_5, x_7, x_8); -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_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_24 = l_Lean_Compiler_LCNF_toLCNFType(x_23, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_24) == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_array_fget_borrowed(x_20, x_17); +x_22 = l_Lean_Expr_fvarId_x21(x_21); +lean_inc_ref(x_6); +x_23 = l_Lean_FVarId_getType___redArg(x_22, x_6, x_8, x_9); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec_ref(x_24); -lean_inc(x_8); -lean_inc_ref(x_7); -x_26 = l_Lean_Compiler_LCNF_toMonoType(x_25, x_7, x_8); -if (lean_obj_tag(x_26) == 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); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_25 = l_Lean_Compiler_LCNF_toLCNFType(x_24, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec_ref(x_26); -lean_inc(x_8); -lean_inc_ref(x_7); -x_28 = l_Lean_IR_toIRType(x_27, x_7, x_8); -if (lean_obj_tag(x_28) == 0) +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec_ref(x_25); +lean_inc(x_9); +lean_inc_ref(x_8); +x_27 = l_Lean_Compiler_LCNF_toMonoType(x_26, x_8, x_9); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_29 = lean_ctor_get(x_4, 1); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); +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); +lean_inc(x_9); +lean_inc_ref(x_8); +x_29 = l_Lean_IR_toIRType(x_28, x_8, x_9); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_30 = lean_ctor_get(x_5, 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_28, 0); +x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); -lean_dec_ref(x_28); -x_34 = lean_ctor_get(x_4, 0); +x_34 = lean_ctor_get(x_29, 0); lean_inc(x_34); -lean_dec_ref(x_4); -x_35 = lean_ctor_get(x_29, 0); +lean_dec_ref(x_29); +x_35 = lean_ctor_get(x_5, 0); lean_inc(x_35); -lean_dec(x_29); +lean_dec_ref(x_5); x_36 = lean_ctor_get(x_30, 0); lean_inc(x_36); lean_dec(x_30); x_37 = lean_ctor_get(x_31, 0); lean_inc(x_37); lean_dec(x_31); -x_38 = !lean_is_exclusive(x_32); -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_32, 0); -x_40 = lean_ctor_get(x_32, 1); -x_41 = lean_unsigned_to_nat(1u); -x_42 = lean_nat_add(x_16, x_41); -lean_dec(x_16); -lean_ctor_set(x_10, 0, x_42); -switch (lean_obj_tag(x_33)) { +x_38 = lean_ctor_get(x_32, 0); +lean_inc(x_38); +lean_dec(x_32); +x_39 = !lean_is_exclusive(x_33); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_33, 0); +x_41 = lean_ctor_get(x_33, 1); +x_42 = lean_unsigned_to_nat(1u); +x_43 = lean_nat_add(x_17, x_42); +lean_dec(x_17); +lean_ctor_set(x_11, 0, x_43); +switch (lean_obj_tag(x_34)) { case 0: { -lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; -lean_free_object(x_32); -lean_dec(x_39); -x_48 = lean_unsigned_to_nat(8u); +lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; +lean_free_object(x_33); +lean_dec(x_40); +x_49 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_49 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_2); -lean_ctor_set(x_49, 2, x_33); -x_50 = lean_unbox(x_35); -lean_dec(x_35); +x_50 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_2); +lean_ctor_set(x_50, 2, x_34); x_51 = lean_unbox(x_36); lean_dec(x_36); x_52 = lean_unbox(x_37); lean_dec(x_37); -x_53 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_50, x_51, x_52, x_17, x_40, x_49, x_5, x_6, x_7, x_8); -x_43 = x_53; -goto block_47; +x_53 = lean_unbox(x_38); +lean_dec(x_38); +x_54 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_51, x_52, x_53, x_18, x_41, x_50, x_6, x_7, x_8, x_9); +x_44 = x_54; +goto block_48; } case 1: { -lean_object* x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; lean_object* x_58; -lean_free_object(x_32); -lean_dec(x_35); -lean_inc(x_2); -x_54 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_54, 0, x_41); -lean_ctor_set(x_54, 1, x_2); -lean_ctor_set(x_54, 2, x_33); -x_55 = lean_unbox(x_36); +lean_object* x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; +lean_free_object(x_33); lean_dec(x_36); +lean_inc(x_2); +x_55 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_55, 0, x_42); +lean_ctor_set(x_55, 1, x_2); +lean_ctor_set(x_55, 2, x_34); x_56 = lean_unbox(x_37); lean_dec(x_37); -x_57 = lean_unbox(x_39); -lean_dec(x_39); -x_58 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_17, x_55, x_56, x_57, x_40, x_54, x_5, x_6, x_7, x_8); -x_43 = x_58; -goto block_47; +x_57 = lean_unbox(x_38); +lean_dec(x_38); +x_58 = lean_unbox(x_40); +lean_dec(x_40); +x_59 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_18, x_56, x_57, x_58, x_41, x_55, x_6, x_7, x_8, x_9); +x_44 = x_59; +goto block_48; } case 2: { -lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; lean_object* x_64; -lean_free_object(x_32); -lean_dec(x_36); -x_59 = lean_unsigned_to_nat(2u); -lean_inc(x_2); -x_60 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_2); -lean_ctor_set(x_60, 2, x_33); -x_61 = lean_unbox(x_35); -lean_dec(x_35); -x_62 = lean_unbox(x_37); +lean_object* x_60; lean_object* x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; lean_object* x_65; +lean_free_object(x_33); lean_dec(x_37); -x_63 = lean_unbox(x_39); -lean_dec(x_39); -x_64 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_61, x_17, x_62, x_63, x_40, x_60, x_5, x_6, x_7, x_8); -x_43 = x_64; -goto block_47; +x_60 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_61 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_2); +lean_ctor_set(x_61, 2, x_34); +x_62 = lean_unbox(x_36); +lean_dec(x_36); +x_63 = lean_unbox(x_38); +lean_dec(x_38); +x_64 = lean_unbox(x_40); +lean_dec(x_40); +x_65 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_62, x_18, x_63, x_64, x_41, x_61, x_6, x_7, x_8, x_9); +x_44 = x_65; +goto block_48; } case 3: { -lean_object* x_65; lean_object* x_66; uint8_t x_67; uint8_t x_68; uint8_t x_69; lean_object* x_70; -lean_free_object(x_32); -lean_dec(x_37); -x_65 = lean_unsigned_to_nat(4u); +lean_object* x_66; lean_object* x_67; uint8_t x_68; uint8_t x_69; uint8_t x_70; lean_object* x_71; +lean_free_object(x_33); +lean_dec(x_38); +x_66 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_66 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_2); -lean_ctor_set(x_66, 2, x_33); -x_67 = lean_unbox(x_35); -lean_dec(x_35); +x_67 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_2); +lean_ctor_set(x_67, 2, x_34); x_68 = lean_unbox(x_36); lean_dec(x_36); -x_69 = lean_unbox(x_39); -lean_dec(x_39); -x_70 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_67, x_68, x_17, x_69, x_40, x_66, x_5, x_6, x_7, x_8); -x_43 = x_70; -goto block_47; +x_69 = lean_unbox(x_37); +lean_dec(x_37); +x_70 = lean_unbox(x_40); +lean_dec(x_40); +x_71 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_68, x_69, x_18, x_70, x_41, x_67, x_6, x_7, x_8, x_9); +x_44 = x_71; +goto block_48; } case 4: { -lean_object* x_71; lean_object* x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; lean_object* x_76; -lean_free_object(x_32); -lean_dec(x_39); -x_71 = lean_unsigned_to_nat(8u); +lean_object* x_72; lean_object* x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; lean_object* x_77; +lean_free_object(x_33); +lean_dec(x_40); +x_72 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_72 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_2); -lean_ctor_set(x_72, 2, x_33); -x_73 = lean_unbox(x_35); -lean_dec(x_35); +x_73 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_2); +lean_ctor_set(x_73, 2, x_34); x_74 = lean_unbox(x_36); lean_dec(x_36); x_75 = lean_unbox(x_37); lean_dec(x_37); -x_76 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_73, x_74, x_75, x_17, x_40, x_72, x_5, x_6, x_7, x_8); -x_43 = x_76; -goto block_47; +x_76 = lean_unbox(x_38); +lean_dec(x_38); +x_77 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_74, x_75, x_76, x_18, x_41, x_73, x_6, x_7, x_8, x_9); +x_44 = x_77; +goto block_48; } case 5: { -lean_object* x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; lean_object* x_82; -lean_free_object(x_32); +lean_object* x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; +lean_free_object(x_33); lean_inc(x_2); -x_77 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_77, 0, x_2); -x_78 = lean_unbox(x_35); -lean_dec(x_35); +x_78 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_78, 0, x_2); +lean_ctor_set_uint8(x_78, sizeof(void*)*1, x_3); x_79 = lean_unbox(x_36); lean_dec(x_36); x_80 = lean_unbox(x_37); lean_dec(x_37); -x_81 = lean_unbox(x_39); -lean_dec(x_39); -x_82 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_78, x_79, x_80, x_81, x_40, x_77, x_5, x_6, x_7, x_8); -x_43 = x_82; -goto block_47; +x_81 = lean_unbox(x_38); +lean_dec(x_38); +x_82 = lean_unbox(x_40); +lean_dec(x_40); +x_83 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_79, x_80, x_81, x_82, x_41, x_78, x_6, x_7, x_8, x_9); +x_44 = x_83; +goto block_48; } case 6: { -lean_object* x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; lean_object* x_88; -lean_free_object(x_32); -x_83 = lean_box(0); -x_84 = lean_unbox(x_35); -lean_dec(x_35); +lean_object* x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; uint8_t x_88; lean_object* x_89; +lean_free_object(x_33); +x_84 = lean_box(0); x_85 = lean_unbox(x_36); lean_dec(x_36); x_86 = lean_unbox(x_37); lean_dec(x_37); -x_87 = lean_unbox(x_39); -lean_dec(x_39); -x_88 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_84, x_85, x_86, x_87, x_40, x_83, x_5, x_6, x_7, x_8); -x_43 = x_88; -goto block_47; +x_87 = lean_unbox(x_38); +lean_dec(x_38); +x_88 = lean_unbox(x_40); +lean_dec(x_40); +x_89 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_85, x_86, x_87, x_88, x_41, x_84, x_6, x_7, x_8, x_9); +x_44 = x_89; +goto block_48; } case 9: { -lean_object* x_89; lean_object* x_90; uint8_t x_91; uint8_t x_92; uint8_t x_93; lean_object* x_94; -lean_free_object(x_32); -lean_dec(x_37); -x_89 = lean_unsigned_to_nat(4u); +lean_object* x_90; lean_object* x_91; uint8_t x_92; uint8_t x_93; uint8_t x_94; lean_object* x_95; +lean_free_object(x_33); +lean_dec(x_38); +x_90 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_90 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_2); -lean_ctor_set(x_90, 2, x_33); -x_91 = lean_unbox(x_35); -lean_dec(x_35); +x_91 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_2); +lean_ctor_set(x_91, 2, x_34); x_92 = lean_unbox(x_36); lean_dec(x_36); -x_93 = lean_unbox(x_39); -lean_dec(x_39); -x_94 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_91, x_92, x_17, x_93, x_40, x_90, x_5, x_6, x_7, x_8); -x_43 = x_94; -goto block_47; +x_93 = lean_unbox(x_37); +lean_dec(x_37); +x_94 = lean_unbox(x_40); +lean_dec(x_40); +x_95 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_92, x_93, x_18, x_94, x_41, x_91, x_6, x_7, x_8, x_9); +x_44 = x_95; +goto block_48; } case 10: { -lean_object* x_95; lean_object* x_96; -lean_free_object(x_32); -lean_dec_ref(x_33); -x_95 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_96 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_95, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_96) == 0) +lean_object* x_96; lean_object* x_97; +lean_free_object(x_33); +lean_dec_ref(x_34); +x_96 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_97 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_96, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_97; uint8_t x_98; uint8_t x_99; uint8_t x_100; uint8_t x_101; lean_object* x_102; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -lean_dec_ref(x_96); -x_98 = lean_unbox(x_35); -lean_dec(x_35); +lean_object* x_98; uint8_t x_99; uint8_t x_100; uint8_t x_101; uint8_t x_102; lean_object* x_103; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +lean_dec_ref(x_97); x_99 = lean_unbox(x_36); lean_dec(x_36); x_100 = lean_unbox(x_37); lean_dec(x_37); -x_101 = lean_unbox(x_39); -lean_dec(x_39); -x_102 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_98, x_99, x_100, x_101, x_40, x_97, x_5, x_6, x_7, x_8); -x_43 = x_102; -goto block_47; +x_101 = lean_unbox(x_38); +lean_dec(x_38); +x_102 = lean_unbox(x_40); +lean_dec(x_40); +x_103 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_99, x_100, x_101, x_102, x_41, x_98, x_6, x_7, x_8, x_9); +x_44 = x_103; +goto block_48; } else { -uint8_t x_103; -lean_dec_ref(x_3); +uint8_t x_104; +lean_dec_ref(x_4); +lean_dec(x_41); lean_dec(x_40); -lean_dec(x_39); +lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_103 = !lean_is_exclusive(x_96); -if (x_103 == 0) +x_104 = !lean_is_exclusive(x_97); +if (x_104 == 0) { -return x_96; +return x_97; } else { -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_96, 0); -lean_inc(x_104); -lean_dec(x_96); -x_105 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_105, 0, x_104); -return x_105; +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_97, 0); +lean_inc(x_105); +lean_dec(x_97); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +return x_106; } } } case 11: { -lean_object* x_106; lean_object* x_107; -lean_free_object(x_32); -lean_dec_ref(x_33); -x_106 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_107 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_106, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_107) == 0) -{ -lean_object* x_108; uint8_t x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; lean_object* x_113; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -lean_dec_ref(x_107); -x_109 = lean_unbox(x_35); -lean_dec(x_35); +lean_object* x_107; lean_object* x_108; +lean_free_object(x_33); +lean_dec_ref(x_34); +x_107 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_108 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_107, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; uint8_t x_113; lean_object* x_114; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +lean_dec_ref(x_108); x_110 = lean_unbox(x_36); lean_dec(x_36); x_111 = lean_unbox(x_37); lean_dec(x_37); -x_112 = lean_unbox(x_39); -lean_dec(x_39); -x_113 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_109, x_110, x_111, x_112, x_40, x_108, x_5, x_6, x_7, x_8); -x_43 = x_113; -goto block_47; +x_112 = lean_unbox(x_38); +lean_dec(x_38); +x_113 = lean_unbox(x_40); +lean_dec(x_40); +x_114 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_110, x_111, x_112, x_113, x_41, x_109, x_6, x_7, x_8, x_9); +x_44 = x_114; +goto block_48; } else { -uint8_t x_114; -lean_dec_ref(x_3); +uint8_t x_115; +lean_dec_ref(x_4); +lean_dec(x_41); lean_dec(x_40); -lean_dec(x_39); +lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_114 = !lean_is_exclusive(x_107); -if (x_114 == 0) +x_115 = !lean_is_exclusive(x_108); +if (x_115 == 0) { -return x_107; +return x_108; } else { -lean_object* x_115; lean_object* x_116; -x_115 = lean_ctor_get(x_107, 0); -lean_inc(x_115); -lean_dec(x_107); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_115); -return x_116; +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_108, 0); +lean_inc(x_116); +lean_dec(x_108); +x_117 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_117, 0, x_116); +return x_117; } } } case 13: { -lean_object* x_117; uint8_t x_118; uint8_t x_119; uint8_t x_120; uint8_t x_121; lean_object* x_122; -lean_free_object(x_32); -x_117 = lean_box(4); -x_118 = lean_unbox(x_35); -lean_dec(x_35); +lean_object* x_118; uint8_t x_119; uint8_t x_120; uint8_t x_121; uint8_t x_122; lean_object* x_123; +lean_free_object(x_33); +x_118 = lean_box(4); x_119 = lean_unbox(x_36); lean_dec(x_36); x_120 = lean_unbox(x_37); lean_dec(x_37); -x_121 = lean_unbox(x_39); -lean_dec(x_39); -x_122 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_118, x_119, x_120, x_121, x_40, x_117, x_5, x_6, x_7, x_8); -x_43 = x_122; -goto block_47; +x_121 = lean_unbox(x_38); +lean_dec(x_38); +x_122 = lean_unbox(x_40); +lean_dec(x_40); +x_123 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_119, x_120, x_121, x_122, x_41, x_118, x_6, x_7, x_8, x_9); +x_44 = x_123; +goto block_48; +} +case 14: +{ +lean_object* x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; lean_object* x_128; +lean_free_object(x_33); +lean_dec(x_36); +lean_inc(x_2); +x_124 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_124, 0, x_42); +lean_ctor_set(x_124, 1, x_2); +lean_ctor_set(x_124, 2, x_34); +x_125 = lean_unbox(x_37); +lean_dec(x_37); +x_126 = lean_unbox(x_38); +lean_dec(x_38); +x_127 = lean_unbox(x_40); +lean_dec(x_40); +x_128 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_18, x_125, x_126, x_127, x_41, x_124, x_6, x_7, x_8, x_9); +x_44 = x_128; +goto block_48; +} +case 15: +{ +lean_object* x_129; lean_object* x_130; uint8_t x_131; uint8_t x_132; uint8_t x_133; lean_object* x_134; +lean_free_object(x_33); +lean_dec(x_37); +x_129 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_130 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_2); +lean_ctor_set(x_130, 2, x_34); +x_131 = lean_unbox(x_36); +lean_dec(x_36); +x_132 = lean_unbox(x_38); +lean_dec(x_38); +x_133 = lean_unbox(x_40); +lean_dec(x_40); +x_134 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_131, x_18, x_132, x_133, x_41, x_130, x_6, x_7, x_8, x_9); +x_44 = x_134; +goto block_48; +} +case 16: +{ +lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; uint8_t x_139; lean_object* x_140; +lean_free_object(x_33); +lean_dec(x_38); +x_135 = lean_unsigned_to_nat(4u); +lean_inc(x_2); +x_136 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_2); +lean_ctor_set(x_136, 2, x_34); +x_137 = lean_unbox(x_36); +lean_dec(x_36); +x_138 = lean_unbox(x_37); +lean_dec(x_37); +x_139 = lean_unbox(x_40); +lean_dec(x_40); +x_140 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_137, x_138, x_18, x_139, x_41, x_136, x_6, x_7, x_8, x_9); +x_44 = x_140; +goto block_48; +} +case 17: +{ +lean_object* x_141; lean_object* x_142; uint8_t x_143; uint8_t x_144; uint8_t x_145; lean_object* x_146; +lean_free_object(x_33); +lean_dec(x_40); +x_141 = lean_unsigned_to_nat(8u); +lean_inc(x_2); +x_142 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_2); +lean_ctor_set(x_142, 2, x_34); +x_143 = lean_unbox(x_36); +lean_dec(x_36); +x_144 = lean_unbox(x_37); +lean_dec(x_37); +x_145 = lean_unbox(x_38); +lean_dec(x_38); +x_146 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_143, x_144, x_145, x_18, x_41, x_142, x_6, x_7, x_8, x_9); +x_44 = x_146; +goto block_48; +} +case 18: +{ +lean_object* x_147; uint8_t x_148; uint8_t x_149; uint8_t x_150; uint8_t x_151; lean_object* x_152; +lean_free_object(x_33); +lean_inc(x_2); +x_147 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_147, 0, x_2); +lean_ctor_set_uint8(x_147, sizeof(void*)*1, x_18); +x_148 = lean_unbox(x_36); +lean_dec(x_36); +x_149 = lean_unbox(x_37); +lean_dec(x_37); +x_150 = lean_unbox(x_38); +lean_dec(x_38); +x_151 = lean_unbox(x_40); +lean_dec(x_40); +x_152 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_148, x_149, x_150, x_151, x_41, x_147, x_6, x_7, x_8, x_9); +x_44 = x_152; +goto block_48; } default: { -lean_object* x_123; uint8_t x_124; uint8_t x_125; uint8_t x_126; uint8_t x_127; lean_object* x_128; -x_123 = lean_nat_add(x_40, x_41); -lean_ctor_set_tag(x_32, 1); -lean_ctor_set(x_32, 1, x_33); -lean_ctor_set(x_32, 0, x_40); -x_124 = lean_unbox(x_35); -lean_dec(x_35); -x_125 = lean_unbox(x_36); +lean_object* x_153; uint8_t x_154; uint8_t x_155; uint8_t x_156; uint8_t x_157; lean_object* x_158; +x_153 = lean_nat_add(x_41, x_42); +lean_ctor_set_tag(x_33, 1); +lean_ctor_set(x_33, 1, x_34); +lean_ctor_set(x_33, 0, x_41); +x_154 = lean_unbox(x_36); lean_dec(x_36); -x_126 = lean_unbox(x_37); +x_155 = lean_unbox(x_37); lean_dec(x_37); -x_127 = lean_unbox(x_39); -lean_dec(x_39); -x_128 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_124, x_125, x_126, x_127, x_123, x_32, x_5, x_6, x_7, x_8); -x_43 = x_128; -goto block_47; +x_156 = lean_unbox(x_38); +lean_dec(x_38); +x_157 = lean_unbox(x_40); +lean_dec(x_40); +x_158 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_154, x_155, x_156, x_157, x_153, x_33, x_6, x_7, x_8, x_9); +x_44 = x_158; +goto block_48; } } -block_47: +block_48: { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -lean_dec_ref(x_43); +lean_object* x_45; lean_object* x_46; x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_44); -x_4 = x_45; +lean_dec_ref(x_44); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +lean_dec(x_45); +x_5 = x_46; goto _start; } } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_129 = lean_ctor_get(x_32, 0); -x_130 = lean_ctor_get(x_32, 1); -lean_inc(x_130); -lean_inc(x_129); -lean_dec(x_32); -x_131 = lean_unsigned_to_nat(1u); -x_132 = lean_nat_add(x_16, x_131); -lean_dec(x_16); -lean_ctor_set(x_10, 0, x_132); -switch (lean_obj_tag(x_33)) { +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_159 = lean_ctor_get(x_33, 0); +x_160 = lean_ctor_get(x_33, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_33); +x_161 = lean_unsigned_to_nat(1u); +x_162 = lean_nat_add(x_17, x_161); +lean_dec(x_17); +lean_ctor_set(x_11, 0, x_162); +switch (lean_obj_tag(x_34)) { case 0: { -lean_object* x_138; lean_object* x_139; uint8_t x_140; uint8_t x_141; uint8_t x_142; lean_object* x_143; -lean_dec(x_129); -x_138 = lean_unsigned_to_nat(8u); +lean_object* x_168; lean_object* x_169; uint8_t x_170; uint8_t x_171; uint8_t x_172; lean_object* x_173; +lean_dec(x_159); +x_168 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_139 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_2); -lean_ctor_set(x_139, 2, x_33); -x_140 = lean_unbox(x_35); -lean_dec(x_35); -x_141 = lean_unbox(x_36); +x_169 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_2); +lean_ctor_set(x_169, 2, x_34); +x_170 = lean_unbox(x_36); lean_dec(x_36); -x_142 = lean_unbox(x_37); +x_171 = lean_unbox(x_37); lean_dec(x_37); -x_143 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_140, x_141, x_142, x_17, x_130, x_139, x_5, x_6, x_7, x_8); -x_133 = x_143; -goto block_137; +x_172 = lean_unbox(x_38); +lean_dec(x_38); +x_173 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_170, x_171, x_172, x_18, x_160, x_169, x_6, x_7, x_8, x_9); +x_163 = x_173; +goto block_167; } case 1: { -lean_object* x_144; uint8_t x_145; uint8_t x_146; uint8_t x_147; lean_object* x_148; -lean_dec(x_35); -lean_inc(x_2); -x_144 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_144, 0, x_131); -lean_ctor_set(x_144, 1, x_2); -lean_ctor_set(x_144, 2, x_33); -x_145 = lean_unbox(x_36); +lean_object* x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; lean_object* x_178; lean_dec(x_36); -x_146 = lean_unbox(x_37); +lean_inc(x_2); +x_174 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_174, 0, x_161); +lean_ctor_set(x_174, 1, x_2); +lean_ctor_set(x_174, 2, x_34); +x_175 = lean_unbox(x_37); lean_dec(x_37); -x_147 = lean_unbox(x_129); -lean_dec(x_129); -x_148 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_17, x_145, x_146, x_147, x_130, x_144, x_5, x_6, x_7, x_8); -x_133 = x_148; -goto block_137; +x_176 = lean_unbox(x_38); +lean_dec(x_38); +x_177 = lean_unbox(x_159); +lean_dec(x_159); +x_178 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_18, x_175, x_176, x_177, x_160, x_174, x_6, x_7, x_8, x_9); +x_163 = x_178; +goto block_167; } case 2: { -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_36); -x_149 = lean_unsigned_to_nat(2u); -lean_inc(x_2); -x_150 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_2); -lean_ctor_set(x_150, 2, x_33); -x_151 = lean_unbox(x_35); -lean_dec(x_35); -x_152 = lean_unbox(x_37); +lean_object* x_179; lean_object* x_180; uint8_t x_181; uint8_t x_182; uint8_t x_183; lean_object* x_184; lean_dec(x_37); -x_153 = lean_unbox(x_129); -lean_dec(x_129); -x_154 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_151, x_17, x_152, x_153, x_130, x_150, x_5, x_6, x_7, x_8); -x_133 = x_154; -goto block_137; +x_179 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_180 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_180, 1, x_2); +lean_ctor_set(x_180, 2, x_34); +x_181 = lean_unbox(x_36); +lean_dec(x_36); +x_182 = lean_unbox(x_38); +lean_dec(x_38); +x_183 = lean_unbox(x_159); +lean_dec(x_159); +x_184 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_181, x_18, x_182, x_183, x_160, x_180, x_6, x_7, x_8, x_9); +x_163 = x_184; +goto block_167; } case 3: { -lean_object* x_155; lean_object* x_156; uint8_t x_157; uint8_t x_158; uint8_t x_159; lean_object* x_160; -lean_dec(x_37); -x_155 = lean_unsigned_to_nat(4u); +lean_object* x_185; lean_object* x_186; uint8_t x_187; uint8_t x_188; uint8_t x_189; lean_object* x_190; +lean_dec(x_38); +x_185 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_156 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_2); -lean_ctor_set(x_156, 2, x_33); -x_157 = lean_unbox(x_35); -lean_dec(x_35); -x_158 = lean_unbox(x_36); +x_186 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_2); +lean_ctor_set(x_186, 2, x_34); +x_187 = lean_unbox(x_36); lean_dec(x_36); -x_159 = lean_unbox(x_129); -lean_dec(x_129); -x_160 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_157, x_158, x_17, x_159, x_130, x_156, x_5, x_6, x_7, x_8); -x_133 = x_160; -goto block_137; +x_188 = lean_unbox(x_37); +lean_dec(x_37); +x_189 = lean_unbox(x_159); +lean_dec(x_159); +x_190 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_187, x_188, x_18, x_189, x_160, x_186, x_6, x_7, x_8, x_9); +x_163 = x_190; +goto block_167; } case 4: { -lean_object* x_161; lean_object* x_162; uint8_t x_163; uint8_t x_164; uint8_t x_165; lean_object* x_166; -lean_dec(x_129); -x_161 = lean_unsigned_to_nat(8u); +lean_object* x_191; lean_object* x_192; uint8_t x_193; uint8_t x_194; uint8_t x_195; lean_object* x_196; +lean_dec(x_159); +x_191 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_162 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_2); -lean_ctor_set(x_162, 2, x_33); -x_163 = lean_unbox(x_35); -lean_dec(x_35); -x_164 = lean_unbox(x_36); +x_192 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_2); +lean_ctor_set(x_192, 2, x_34); +x_193 = lean_unbox(x_36); lean_dec(x_36); -x_165 = lean_unbox(x_37); +x_194 = lean_unbox(x_37); lean_dec(x_37); -x_166 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_163, x_164, x_165, x_17, x_130, x_162, x_5, x_6, x_7, x_8); -x_133 = x_166; -goto block_137; +x_195 = lean_unbox(x_38); +lean_dec(x_38); +x_196 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_193, x_194, x_195, x_18, x_160, x_192, x_6, x_7, x_8, x_9); +x_163 = x_196; +goto block_167; } case 5: { -lean_object* x_167; uint8_t x_168; uint8_t x_169; uint8_t x_170; uint8_t x_171; lean_object* x_172; +lean_object* x_197; uint8_t x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; lean_object* x_202; lean_inc(x_2); -x_167 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_167, 0, x_2); -x_168 = lean_unbox(x_35); -lean_dec(x_35); -x_169 = lean_unbox(x_36); +x_197 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_197, 0, x_2); +lean_ctor_set_uint8(x_197, sizeof(void*)*1, x_3); +x_198 = lean_unbox(x_36); lean_dec(x_36); -x_170 = lean_unbox(x_37); +x_199 = lean_unbox(x_37); lean_dec(x_37); -x_171 = lean_unbox(x_129); -lean_dec(x_129); -x_172 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_168, x_169, x_170, x_171, x_130, x_167, x_5, x_6, x_7, x_8); -x_133 = x_172; -goto block_137; +x_200 = lean_unbox(x_38); +lean_dec(x_38); +x_201 = lean_unbox(x_159); +lean_dec(x_159); +x_202 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_198, x_199, x_200, x_201, x_160, x_197, x_6, x_7, x_8, x_9); +x_163 = x_202; +goto block_167; } case 6: { -lean_object* x_173; uint8_t x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; lean_object* x_178; -x_173 = lean_box(0); -x_174 = lean_unbox(x_35); -lean_dec(x_35); -x_175 = lean_unbox(x_36); +lean_object* x_203; uint8_t x_204; uint8_t x_205; uint8_t x_206; uint8_t x_207; lean_object* x_208; +x_203 = lean_box(0); +x_204 = lean_unbox(x_36); lean_dec(x_36); -x_176 = lean_unbox(x_37); +x_205 = lean_unbox(x_37); lean_dec(x_37); -x_177 = lean_unbox(x_129); -lean_dec(x_129); -x_178 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_174, x_175, x_176, x_177, x_130, x_173, x_5, x_6, x_7, x_8); -x_133 = x_178; -goto block_137; +x_206 = lean_unbox(x_38); +lean_dec(x_38); +x_207 = lean_unbox(x_159); +lean_dec(x_159); +x_208 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_204, x_205, x_206, x_207, x_160, x_203, x_6, x_7, x_8, x_9); +x_163 = x_208; +goto block_167; } case 9: { -lean_object* x_179; lean_object* x_180; uint8_t x_181; uint8_t x_182; uint8_t x_183; lean_object* x_184; -lean_dec(x_37); -x_179 = lean_unsigned_to_nat(4u); +lean_object* x_209; lean_object* x_210; uint8_t x_211; uint8_t x_212; uint8_t x_213; lean_object* x_214; +lean_dec(x_38); +x_209 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_180 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_180, 0, x_179); -lean_ctor_set(x_180, 1, x_2); -lean_ctor_set(x_180, 2, x_33); -x_181 = lean_unbox(x_35); -lean_dec(x_35); -x_182 = lean_unbox(x_36); +x_210 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_210, 0, x_209); +lean_ctor_set(x_210, 1, x_2); +lean_ctor_set(x_210, 2, x_34); +x_211 = lean_unbox(x_36); lean_dec(x_36); -x_183 = lean_unbox(x_129); -lean_dec(x_129); -x_184 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_181, x_182, x_17, x_183, x_130, x_180, x_5, x_6, x_7, x_8); -x_133 = x_184; -goto block_137; +x_212 = lean_unbox(x_37); +lean_dec(x_37); +x_213 = lean_unbox(x_159); +lean_dec(x_159); +x_214 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_211, x_212, x_18, x_213, x_160, x_210, x_6, x_7, x_8, x_9); +x_163 = x_214; +goto block_167; } case 10: { -lean_object* x_185; lean_object* x_186; -lean_dec_ref(x_33); -x_185 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_186 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_185, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_186) == 0) -{ -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_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -lean_dec_ref(x_186); -x_188 = lean_unbox(x_35); -lean_dec(x_35); -x_189 = lean_unbox(x_36); +lean_object* x_215; lean_object* x_216; +lean_dec_ref(x_34); +x_215 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_216 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_215, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; uint8_t x_218; uint8_t x_219; uint8_t x_220; uint8_t x_221; lean_object* x_222; +x_217 = lean_ctor_get(x_216, 0); +lean_inc(x_217); +lean_dec_ref(x_216); +x_218 = lean_unbox(x_36); lean_dec(x_36); -x_190 = lean_unbox(x_37); +x_219 = lean_unbox(x_37); lean_dec(x_37); -x_191 = lean_unbox(x_129); -lean_dec(x_129); -x_192 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_188, x_189, x_190, x_191, x_130, x_187, x_5, x_6, x_7, x_8); -x_133 = x_192; -goto block_137; +x_220 = lean_unbox(x_38); +lean_dec(x_38); +x_221 = lean_unbox(x_159); +lean_dec(x_159); +x_222 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_218, x_219, x_220, x_221, x_160, x_217, x_6, x_7, x_8, x_9); +x_163 = x_222; +goto block_167; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec_ref(x_3); -lean_dec(x_130); -lean_dec(x_129); +lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec_ref(x_4); +lean_dec(x_160); +lean_dec(x_159); +lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_193 = lean_ctor_get(x_186, 0); -lean_inc(x_193); -if (lean_is_exclusive(x_186)) { - lean_ctor_release(x_186, 0); - x_194 = x_186; +x_223 = lean_ctor_get(x_216, 0); +lean_inc(x_223); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + x_224 = x_216; } else { - lean_dec_ref(x_186); - x_194 = lean_box(0); + lean_dec_ref(x_216); + x_224 = lean_box(0); } -if (lean_is_scalar(x_194)) { - x_195 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_224)) { + x_225 = lean_alloc_ctor(1, 1, 0); } else { - x_195 = x_194; + x_225 = x_224; } -lean_ctor_set(x_195, 0, x_193); -return x_195; +lean_ctor_set(x_225, 0, x_223); +return x_225; } } case 11: { -lean_object* x_196; lean_object* x_197; -lean_dec_ref(x_33); -x_196 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_197 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_196, x_5, x_6, x_7, x_8); -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_35); -lean_dec(x_35); -x_200 = lean_unbox(x_36); +lean_object* x_226; lean_object* x_227; +lean_dec_ref(x_34); +x_226 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_227 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_226, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_227) == 0) +{ +lean_object* x_228; uint8_t x_229; uint8_t x_230; uint8_t x_231; uint8_t x_232; lean_object* x_233; +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +lean_dec_ref(x_227); +x_229 = lean_unbox(x_36); lean_dec(x_36); -x_201 = lean_unbox(x_37); +x_230 = lean_unbox(x_37); lean_dec(x_37); -x_202 = lean_unbox(x_129); -lean_dec(x_129); -x_203 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_199, x_200, x_201, x_202, x_130, x_198, x_5, x_6, x_7, x_8); -x_133 = x_203; -goto block_137; +x_231 = lean_unbox(x_38); +lean_dec(x_38); +x_232 = lean_unbox(x_159); +lean_dec(x_159); +x_233 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_229, x_230, x_231, x_232, x_160, x_228, x_6, x_7, x_8, x_9); +x_163 = x_233; +goto block_167; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec_ref(x_3); -lean_dec(x_130); -lean_dec(x_129); +lean_object* x_234; lean_object* x_235; lean_object* x_236; +lean_dec_ref(x_4); +lean_dec(x_160); +lean_dec(x_159); +lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -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; +x_234 = lean_ctor_get(x_227, 0); +lean_inc(x_234); +if (lean_is_exclusive(x_227)) { + lean_ctor_release(x_227, 0); + x_235 = x_227; } else { - lean_dec_ref(x_197); - x_205 = lean_box(0); + lean_dec_ref(x_227); + x_235 = lean_box(0); } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_235)) { + x_236 = lean_alloc_ctor(1, 1, 0); } else { - x_206 = x_205; + x_236 = x_235; } -lean_ctor_set(x_206, 0, x_204); -return x_206; +lean_ctor_set(x_236, 0, x_234); +return x_236; } } case 13: { -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_207 = lean_box(4); -x_208 = lean_unbox(x_35); -lean_dec(x_35); -x_209 = lean_unbox(x_36); +lean_object* x_237; uint8_t x_238; uint8_t x_239; uint8_t x_240; uint8_t x_241; lean_object* x_242; +x_237 = lean_box(4); +x_238 = lean_unbox(x_36); +lean_dec(x_36); +x_239 = lean_unbox(x_37); +lean_dec(x_37); +x_240 = lean_unbox(x_38); +lean_dec(x_38); +x_241 = lean_unbox(x_159); +lean_dec(x_159); +x_242 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_238, x_239, x_240, x_241, x_160, x_237, x_6, x_7, x_8, x_9); +x_163 = x_242; +goto block_167; +} +case 14: +{ +lean_object* x_243; uint8_t x_244; uint8_t x_245; uint8_t x_246; lean_object* x_247; +lean_dec(x_36); +lean_inc(x_2); +x_243 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_243, 0, x_161); +lean_ctor_set(x_243, 1, x_2); +lean_ctor_set(x_243, 2, x_34); +x_244 = lean_unbox(x_37); +lean_dec(x_37); +x_245 = lean_unbox(x_38); +lean_dec(x_38); +x_246 = lean_unbox(x_159); +lean_dec(x_159); +x_247 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_18, x_244, x_245, x_246, x_160, x_243, x_6, x_7, x_8, x_9); +x_163 = x_247; +goto block_167; +} +case 15: +{ +lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; uint8_t x_252; lean_object* x_253; +lean_dec(x_37); +x_248 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_249 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_2); +lean_ctor_set(x_249, 2, x_34); +x_250 = lean_unbox(x_36); +lean_dec(x_36); +x_251 = lean_unbox(x_38); +lean_dec(x_38); +x_252 = lean_unbox(x_159); +lean_dec(x_159); +x_253 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_250, x_18, x_251, x_252, x_160, x_249, x_6, x_7, x_8, x_9); +x_163 = x_253; +goto block_167; +} +case 16: +{ +lean_object* x_254; lean_object* x_255; uint8_t x_256; uint8_t x_257; uint8_t x_258; lean_object* x_259; +lean_dec(x_38); +x_254 = lean_unsigned_to_nat(4u); +lean_inc(x_2); +x_255 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_2); +lean_ctor_set(x_255, 2, x_34); +x_256 = lean_unbox(x_36); lean_dec(x_36); -x_210 = lean_unbox(x_37); +x_257 = lean_unbox(x_37); lean_dec(x_37); -x_211 = lean_unbox(x_129); -lean_dec(x_129); -x_212 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_208, x_209, x_210, x_211, x_130, x_207, x_5, x_6, x_7, x_8); -x_133 = x_212; -goto block_137; +x_258 = lean_unbox(x_159); +lean_dec(x_159); +x_259 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_256, x_257, x_18, x_258, x_160, x_255, x_6, x_7, x_8, x_9); +x_163 = x_259; +goto block_167; +} +case 17: +{ +lean_object* x_260; lean_object* x_261; uint8_t x_262; uint8_t x_263; uint8_t x_264; lean_object* x_265; +lean_dec(x_159); +x_260 = lean_unsigned_to_nat(8u); +lean_inc(x_2); +x_261 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_2); +lean_ctor_set(x_261, 2, x_34); +x_262 = lean_unbox(x_36); +lean_dec(x_36); +x_263 = lean_unbox(x_37); +lean_dec(x_37); +x_264 = lean_unbox(x_38); +lean_dec(x_38); +x_265 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_262, x_263, x_264, x_18, x_160, x_261, x_6, x_7, x_8, x_9); +x_163 = x_265; +goto block_167; +} +case 18: +{ +lean_object* x_266; uint8_t x_267; uint8_t x_268; uint8_t x_269; uint8_t x_270; lean_object* x_271; +lean_inc(x_2); +x_266 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_266, 0, x_2); +lean_ctor_set_uint8(x_266, sizeof(void*)*1, x_18); +x_267 = lean_unbox(x_36); +lean_dec(x_36); +x_268 = lean_unbox(x_37); +lean_dec(x_37); +x_269 = lean_unbox(x_38); +lean_dec(x_38); +x_270 = lean_unbox(x_159); +lean_dec(x_159); +x_271 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_267, x_268, x_269, x_270, x_160, x_266, x_6, x_7, x_8, x_9); +x_163 = x_271; +goto block_167; } default: { -lean_object* x_213; lean_object* x_214; uint8_t x_215; uint8_t x_216; uint8_t x_217; uint8_t x_218; lean_object* x_219; -x_213 = lean_nat_add(x_130, x_131); -x_214 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_214, 0, x_130); -lean_ctor_set(x_214, 1, x_33); -x_215 = lean_unbox(x_35); -lean_dec(x_35); -x_216 = lean_unbox(x_36); +lean_object* x_272; lean_object* x_273; uint8_t x_274; uint8_t x_275; uint8_t x_276; uint8_t x_277; lean_object* x_278; +x_272 = lean_nat_add(x_160, x_161); +x_273 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_273, 0, x_160); +lean_ctor_set(x_273, 1, x_34); +x_274 = lean_unbox(x_36); lean_dec(x_36); -x_217 = lean_unbox(x_37); +x_275 = lean_unbox(x_37); lean_dec(x_37); -x_218 = lean_unbox(x_129); -lean_dec(x_129); -x_219 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_34, x_215, x_216, x_217, x_218, x_213, x_214, x_5, x_6, x_7, x_8); -x_133 = x_219; -goto block_137; -} -} -block_137: -{ -lean_object* x_134; lean_object* x_135; -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -lean_dec_ref(x_133); -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -lean_dec(x_134); -x_4 = x_135; +x_276 = lean_unbox(x_38); +lean_dec(x_38); +x_277 = lean_unbox(x_159); +lean_dec(x_159); +x_278 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_35, x_274, x_275, x_276, x_277, x_272, x_273, x_6, x_7, x_8, x_9); +x_163 = x_278; +goto block_167; +} +} +block_167: +{ +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +lean_dec_ref(x_163); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +lean_dec(x_164); +x_5 = x_165; goto _start; } } } else { -uint8_t x_220; -lean_free_object(x_10); -lean_dec(x_16); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +uint8_t x_279; +lean_free_object(x_11); +lean_dec(x_17); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_220 = !lean_is_exclusive(x_28); -if (x_220 == 0) +x_279 = !lean_is_exclusive(x_29); +if (x_279 == 0) { -return x_28; +return x_29; } else { -lean_object* x_221; lean_object* x_222; -x_221 = lean_ctor_get(x_28, 0); -lean_inc(x_221); -lean_dec(x_28); -x_222 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_222, 0, x_221); -return x_222; +lean_object* x_280; lean_object* x_281; +x_280 = lean_ctor_get(x_29, 0); +lean_inc(x_280); +lean_dec(x_29); +x_281 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_281, 0, x_280); +return x_281; } } } else { -uint8_t x_223; -lean_free_object(x_10); -lean_dec(x_16); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +uint8_t x_282; +lean_free_object(x_11); +lean_dec(x_17); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_223 = !lean_is_exclusive(x_26); -if (x_223 == 0) +x_282 = !lean_is_exclusive(x_27); +if (x_282 == 0) { -return x_26; +return x_27; } else { -lean_object* x_224; lean_object* x_225; -x_224 = lean_ctor_get(x_26, 0); -lean_inc(x_224); -lean_dec(x_26); -x_225 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_225, 0, x_224); -return x_225; +lean_object* x_283; lean_object* x_284; +x_283 = lean_ctor_get(x_27, 0); +lean_inc(x_283); +lean_dec(x_27); +x_284 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_284, 0, x_283); +return x_284; } } } else { -uint8_t x_226; -lean_free_object(x_10); -lean_dec(x_16); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +uint8_t x_285; +lean_free_object(x_11); +lean_dec(x_17); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_226 = !lean_is_exclusive(x_24); -if (x_226 == 0) +x_285 = !lean_is_exclusive(x_25); +if (x_285 == 0) { -return x_24; +return x_25; } else { -lean_object* x_227; lean_object* x_228; -x_227 = lean_ctor_get(x_24, 0); -lean_inc(x_227); -lean_dec(x_24); -x_228 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_228, 0, x_227); -return x_228; +lean_object* x_286; lean_object* x_287; +x_286 = lean_ctor_get(x_25, 0); +lean_inc(x_286); +lean_dec(x_25); +x_287 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_287, 0, x_286); +return x_287; } } } else { -uint8_t x_229; -lean_free_object(x_10); -lean_dec(x_16); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +uint8_t x_288; +lean_free_object(x_11); +lean_dec(x_17); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_229 = !lean_is_exclusive(x_22); -if (x_229 == 0) +x_288 = !lean_is_exclusive(x_23); +if (x_288 == 0) { -return x_22; +return x_23; } else { -lean_object* x_230; lean_object* x_231; -x_230 = lean_ctor_get(x_22, 0); -lean_inc(x_230); -lean_dec(x_22); -x_231 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_231, 0, x_230); -return x_231; +lean_object* x_289; lean_object* x_290; +x_289 = lean_ctor_get(x_23, 0); +lean_inc(x_289); +lean_dec(x_23); +x_290 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_290, 0, x_289); +return x_290; } } } } else { -lean_object* x_232; uint8_t x_233; -x_232 = lean_ctor_get(x_10, 0); -lean_inc(x_232); -lean_dec(x_10); -x_233 = lean_nat_dec_lt(x_232, x_13); -if (x_233 == 0) +lean_object* x_291; uint8_t x_292; +x_291 = lean_ctor_get(x_11, 0); +lean_inc(x_291); +lean_dec(x_11); +x_292 = lean_nat_dec_lt(x_291, x_14); +if (x_292 == 0) { -lean_object* x_234; -lean_dec(x_232); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_293; +lean_dec(x_291); +lean_free_object(x_4); +lean_dec(x_14); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_234 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_234, 0, x_4); -return x_234; +x_293 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_293, 0, x_5); +return x_293; } else { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_235 = lean_ctor_get(x_1, 0); -x_236 = lean_array_fget_borrowed(x_235, x_232); -x_237 = l_Lean_Expr_fvarId_x21(x_236); -lean_inc_ref(x_5); -x_238 = l_Lean_FVarId_getType___redArg(x_237, x_5, x_7, x_8); -if (lean_obj_tag(x_238) == 0) +lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_294 = lean_ctor_get(x_1, 0); +x_295 = lean_array_fget_borrowed(x_294, x_291); +x_296 = l_Lean_Expr_fvarId_x21(x_295); +lean_inc_ref(x_6); +x_297 = l_Lean_FVarId_getType___redArg(x_296, x_6, x_8, x_9); +if (lean_obj_tag(x_297) == 0) { -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_238, 0); -lean_inc(x_239); -lean_dec_ref(x_238); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_240 = l_Lean_Compiler_LCNF_toLCNFType(x_239, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_240) == 0) -{ -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_240, 0); -lean_inc(x_241); -lean_dec_ref(x_240); -lean_inc(x_8); -lean_inc_ref(x_7); -x_242 = l_Lean_Compiler_LCNF_toMonoType(x_241, x_7, x_8); -if (lean_obj_tag(x_242) == 0) -{ -lean_object* x_243; lean_object* x_244; -x_243 = lean_ctor_get(x_242, 0); -lean_inc(x_243); -lean_dec_ref(x_242); -lean_inc(x_8); -lean_inc_ref(x_7); -x_244 = l_Lean_IR_toIRType(x_243, x_7, x_8); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; 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_245 = lean_ctor_get(x_4, 1); -lean_inc(x_245); -x_246 = lean_ctor_get(x_245, 1); -lean_inc(x_246); -x_247 = lean_ctor_get(x_246, 1); -lean_inc(x_247); -x_248 = lean_ctor_get(x_247, 1); -lean_inc(x_248); -x_249 = lean_ctor_get(x_244, 0); -lean_inc(x_249); -lean_dec_ref(x_244); -x_250 = lean_ctor_get(x_4, 0); -lean_inc(x_250); -lean_dec_ref(x_4); -x_251 = lean_ctor_get(x_245, 0); -lean_inc(x_251); -lean_dec(x_245); -x_252 = lean_ctor_get(x_246, 0); -lean_inc(x_252); -lean_dec(x_246); -x_253 = lean_ctor_get(x_247, 0); -lean_inc(x_253); -lean_dec(x_247); -x_254 = lean_ctor_get(x_248, 0); -lean_inc(x_254); -x_255 = lean_ctor_get(x_248, 1); -lean_inc(x_255); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_256 = x_248; +lean_object* x_298; lean_object* x_299; +x_298 = lean_ctor_get(x_297, 0); +lean_inc(x_298); +lean_dec_ref(x_297); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_299 = l_Lean_Compiler_LCNF_toLCNFType(x_298, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_299) == 0) +{ +lean_object* x_300; lean_object* x_301; +x_300 = lean_ctor_get(x_299, 0); +lean_inc(x_300); +lean_dec_ref(x_299); +lean_inc(x_9); +lean_inc_ref(x_8); +x_301 = l_Lean_Compiler_LCNF_toMonoType(x_300, x_8, x_9); +if (lean_obj_tag(x_301) == 0) +{ +lean_object* x_302; lean_object* x_303; +x_302 = lean_ctor_get(x_301, 0); +lean_inc(x_302); +lean_dec_ref(x_301); +lean_inc(x_9); +lean_inc_ref(x_8); +x_303 = l_Lean_IR_toIRType(x_302, x_8, x_9); +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; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_304 = lean_ctor_get(x_5, 1); +lean_inc(x_304); +x_305 = lean_ctor_get(x_304, 1); +lean_inc(x_305); +x_306 = lean_ctor_get(x_305, 1); +lean_inc(x_306); +x_307 = lean_ctor_get(x_306, 1); +lean_inc(x_307); +x_308 = lean_ctor_get(x_303, 0); +lean_inc(x_308); +lean_dec_ref(x_303); +x_309 = lean_ctor_get(x_5, 0); +lean_inc(x_309); +lean_dec_ref(x_5); +x_310 = lean_ctor_get(x_304, 0); +lean_inc(x_310); +lean_dec(x_304); +x_311 = lean_ctor_get(x_305, 0); +lean_inc(x_311); +lean_dec(x_305); +x_312 = lean_ctor_get(x_306, 0); +lean_inc(x_312); +lean_dec(x_306); +x_313 = lean_ctor_get(x_307, 0); +lean_inc(x_313); +x_314 = lean_ctor_get(x_307, 1); +lean_inc(x_314); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_315 = x_307; } else { - lean_dec_ref(x_248); - x_256 = lean_box(0); -} -x_257 = lean_unsigned_to_nat(1u); -x_258 = lean_nat_add(x_232, x_257); -lean_dec(x_232); -x_259 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_3, 0, x_259); -switch (lean_obj_tag(x_249)) { + lean_dec_ref(x_307); + x_315 = lean_box(0); +} +x_316 = lean_unsigned_to_nat(1u); +x_317 = lean_nat_add(x_291, x_316); +lean_dec(x_291); +x_318 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_4, 0, x_318); +switch (lean_obj_tag(x_308)) { 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_256); -lean_dec(x_254); -x_265 = lean_unsigned_to_nat(8u); +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_315); +lean_dec(x_313); +x_324 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_266 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_266, 0, x_265); -lean_ctor_set(x_266, 1, x_2); -lean_ctor_set(x_266, 2, x_249); -x_267 = lean_unbox(x_251); -lean_dec(x_251); -x_268 = lean_unbox(x_252); -lean_dec(x_252); -x_269 = lean_unbox(x_253); -lean_dec(x_253); -x_270 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_267, x_268, x_269, x_233, x_255, x_266, x_5, x_6, x_7, x_8); -x_260 = x_270; -goto block_264; +x_325 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_325, 0, x_324); +lean_ctor_set(x_325, 1, x_2); +lean_ctor_set(x_325, 2, x_308); +x_326 = lean_unbox(x_310); +lean_dec(x_310); +x_327 = lean_unbox(x_311); +lean_dec(x_311); +x_328 = lean_unbox(x_312); +lean_dec(x_312); +x_329 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_326, x_327, x_328, x_292, x_314, x_325, x_6, x_7, x_8, x_9); +x_319 = x_329; +goto block_323; } 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_256); -lean_dec(x_251); +lean_object* x_330; uint8_t x_331; uint8_t x_332; uint8_t x_333; lean_object* x_334; +lean_dec(x_315); +lean_dec(x_310); lean_inc(x_2); -x_271 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_271, 0, x_257); -lean_ctor_set(x_271, 1, x_2); -lean_ctor_set(x_271, 2, x_249); -x_272 = lean_unbox(x_252); -lean_dec(x_252); -x_273 = lean_unbox(x_253); -lean_dec(x_253); -x_274 = lean_unbox(x_254); -lean_dec(x_254); -x_275 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_233, x_272, x_273, x_274, x_255, x_271, x_5, x_6, x_7, x_8); -x_260 = x_275; -goto block_264; +x_330 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_330, 0, x_316); +lean_ctor_set(x_330, 1, x_2); +lean_ctor_set(x_330, 2, x_308); +x_331 = lean_unbox(x_311); +lean_dec(x_311); +x_332 = lean_unbox(x_312); +lean_dec(x_312); +x_333 = lean_unbox(x_313); +lean_dec(x_313); +x_334 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_292, x_331, x_332, x_333, x_314, x_330, x_6, x_7, x_8, x_9); +x_319 = x_334; +goto block_323; } 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_256); -lean_dec(x_252); -x_276 = lean_unsigned_to_nat(2u); +lean_object* x_335; lean_object* x_336; uint8_t x_337; uint8_t x_338; uint8_t x_339; lean_object* x_340; +lean_dec(x_315); +lean_dec(x_311); +x_335 = lean_unsigned_to_nat(2u); lean_inc(x_2); -x_277 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_277, 0, x_276); -lean_ctor_set(x_277, 1, x_2); -lean_ctor_set(x_277, 2, x_249); -x_278 = lean_unbox(x_251); -lean_dec(x_251); -x_279 = lean_unbox(x_253); -lean_dec(x_253); -x_280 = lean_unbox(x_254); -lean_dec(x_254); -x_281 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_278, x_233, x_279, x_280, x_255, x_277, x_5, x_6, x_7, x_8); -x_260 = x_281; -goto block_264; +x_336 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_336, 0, x_335); +lean_ctor_set(x_336, 1, x_2); +lean_ctor_set(x_336, 2, x_308); +x_337 = lean_unbox(x_310); +lean_dec(x_310); +x_338 = lean_unbox(x_312); +lean_dec(x_312); +x_339 = lean_unbox(x_313); +lean_dec(x_313); +x_340 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_337, x_292, x_338, x_339, x_314, x_336, x_6, x_7, x_8, x_9); +x_319 = x_340; +goto block_323; } 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_256); -lean_dec(x_253); -x_282 = lean_unsigned_to_nat(4u); +lean_object* x_341; lean_object* x_342; uint8_t x_343; uint8_t x_344; uint8_t x_345; lean_object* x_346; +lean_dec(x_315); +lean_dec(x_312); +x_341 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_283 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_283, 0, x_282); -lean_ctor_set(x_283, 1, x_2); -lean_ctor_set(x_283, 2, x_249); -x_284 = lean_unbox(x_251); -lean_dec(x_251); -x_285 = lean_unbox(x_252); -lean_dec(x_252); -x_286 = lean_unbox(x_254); -lean_dec(x_254); -x_287 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_284, x_285, x_233, x_286, x_255, x_283, x_5, x_6, x_7, x_8); -x_260 = x_287; -goto block_264; +x_342 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_342, 0, x_341); +lean_ctor_set(x_342, 1, x_2); +lean_ctor_set(x_342, 2, x_308); +x_343 = lean_unbox(x_310); +lean_dec(x_310); +x_344 = lean_unbox(x_311); +lean_dec(x_311); +x_345 = lean_unbox(x_313); +lean_dec(x_313); +x_346 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_343, x_344, x_292, x_345, x_314, x_342, x_6, x_7, x_8, x_9); +x_319 = x_346; +goto block_323; } 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_256); -lean_dec(x_254); -x_288 = lean_unsigned_to_nat(8u); +lean_object* x_347; lean_object* x_348; uint8_t x_349; uint8_t x_350; uint8_t x_351; lean_object* x_352; +lean_dec(x_315); +lean_dec(x_313); +x_347 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_289 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_289, 0, x_288); -lean_ctor_set(x_289, 1, x_2); -lean_ctor_set(x_289, 2, x_249); -x_290 = lean_unbox(x_251); -lean_dec(x_251); -x_291 = lean_unbox(x_252); -lean_dec(x_252); -x_292 = lean_unbox(x_253); -lean_dec(x_253); -x_293 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_290, x_291, x_292, x_233, x_255, x_289, x_5, x_6, x_7, x_8); -x_260 = x_293; -goto block_264; +x_348 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_348, 0, x_347); +lean_ctor_set(x_348, 1, x_2); +lean_ctor_set(x_348, 2, x_308); +x_349 = lean_unbox(x_310); +lean_dec(x_310); +x_350 = lean_unbox(x_311); +lean_dec(x_311); +x_351 = lean_unbox(x_312); +lean_dec(x_312); +x_352 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_349, x_350, x_351, x_292, x_314, x_348, x_6, x_7, x_8, x_9); +x_319 = x_352; +goto block_323; } 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_256); +lean_object* x_353; uint8_t x_354; uint8_t x_355; uint8_t x_356; uint8_t x_357; lean_object* x_358; +lean_dec(x_315); lean_inc(x_2); -x_294 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_294, 0, x_2); -x_295 = lean_unbox(x_251); -lean_dec(x_251); -x_296 = lean_unbox(x_252); -lean_dec(x_252); -x_297 = lean_unbox(x_253); -lean_dec(x_253); -x_298 = lean_unbox(x_254); -lean_dec(x_254); -x_299 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_295, x_296, x_297, x_298, x_255, x_294, x_5, x_6, x_7, x_8); -x_260 = x_299; -goto block_264; +x_353 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_353, 0, x_2); +lean_ctor_set_uint8(x_353, sizeof(void*)*1, x_3); +x_354 = lean_unbox(x_310); +lean_dec(x_310); +x_355 = lean_unbox(x_311); +lean_dec(x_311); +x_356 = lean_unbox(x_312); +lean_dec(x_312); +x_357 = lean_unbox(x_313); +lean_dec(x_313); +x_358 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_354, x_355, x_356, x_357, x_314, x_353, x_6, x_7, x_8, x_9); +x_319 = x_358; +goto block_323; } 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_256); -x_300 = lean_box(0); -x_301 = lean_unbox(x_251); -lean_dec(x_251); -x_302 = lean_unbox(x_252); -lean_dec(x_252); -x_303 = lean_unbox(x_253); -lean_dec(x_253); -x_304 = lean_unbox(x_254); -lean_dec(x_254); -x_305 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_301, x_302, x_303, x_304, x_255, x_300, x_5, x_6, x_7, x_8); -x_260 = x_305; -goto block_264; +lean_object* x_359; uint8_t x_360; uint8_t x_361; uint8_t x_362; uint8_t x_363; lean_object* x_364; +lean_dec(x_315); +x_359 = lean_box(0); +x_360 = lean_unbox(x_310); +lean_dec(x_310); +x_361 = lean_unbox(x_311); +lean_dec(x_311); +x_362 = lean_unbox(x_312); +lean_dec(x_312); +x_363 = lean_unbox(x_313); +lean_dec(x_313); +x_364 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_360, x_361, x_362, x_363, x_314, x_359, x_6, x_7, x_8, x_9); +x_319 = x_364; +goto block_323; } 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_256); -lean_dec(x_253); -x_306 = lean_unsigned_to_nat(4u); +lean_object* x_365; lean_object* x_366; uint8_t x_367; uint8_t x_368; uint8_t x_369; lean_object* x_370; +lean_dec(x_315); +lean_dec(x_312); +x_365 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_307 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_2); -lean_ctor_set(x_307, 2, x_249); -x_308 = lean_unbox(x_251); -lean_dec(x_251); -x_309 = lean_unbox(x_252); -lean_dec(x_252); -x_310 = lean_unbox(x_254); -lean_dec(x_254); -x_311 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_308, x_309, x_233, x_310, x_255, x_307, x_5, x_6, x_7, x_8); -x_260 = x_311; -goto block_264; +x_366 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_2); +lean_ctor_set(x_366, 2, x_308); +x_367 = lean_unbox(x_310); +lean_dec(x_310); +x_368 = lean_unbox(x_311); +lean_dec(x_311); +x_369 = lean_unbox(x_313); +lean_dec(x_313); +x_370 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_367, x_368, x_292, x_369, x_314, x_366, x_6, x_7, x_8, x_9); +x_319 = x_370; +goto block_323; } case 10: { -lean_object* x_312; lean_object* x_313; -lean_dec(x_256); -lean_dec_ref(x_249); -x_312 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_313 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_312, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; uint8_t x_315; uint8_t x_316; uint8_t x_317; uint8_t x_318; lean_object* x_319; -x_314 = lean_ctor_get(x_313, 0); -lean_inc(x_314); -lean_dec_ref(x_313); -x_315 = lean_unbox(x_251); -lean_dec(x_251); -x_316 = lean_unbox(x_252); -lean_dec(x_252); -x_317 = lean_unbox(x_253); -lean_dec(x_253); -x_318 = lean_unbox(x_254); -lean_dec(x_254); -x_319 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_315, x_316, x_317, x_318, x_255, x_314, x_5, x_6, x_7, x_8); -x_260 = x_319; -goto block_264; -} -else -{ -lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec_ref(x_3); -lean_dec(x_255); -lean_dec(x_254); -lean_dec(x_253); -lean_dec(x_252); -lean_dec(x_251); -lean_dec(x_250); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_371; lean_object* x_372; +lean_dec(x_315); +lean_dec_ref(x_308); +x_371 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_372 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_371, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_372) == 0) +{ +lean_object* x_373; uint8_t x_374; uint8_t x_375; uint8_t x_376; uint8_t x_377; lean_object* x_378; +x_373 = lean_ctor_get(x_372, 0); +lean_inc(x_373); +lean_dec_ref(x_372); +x_374 = lean_unbox(x_310); +lean_dec(x_310); +x_375 = lean_unbox(x_311); +lean_dec(x_311); +x_376 = lean_unbox(x_312); +lean_dec(x_312); +x_377 = lean_unbox(x_313); +lean_dec(x_313); +x_378 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_374, x_375, x_376, x_377, x_314, x_373, x_6, x_7, x_8, x_9); +x_319 = x_378; +goto block_323; +} +else +{ +lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec_ref(x_4); +lean_dec(x_314); +lean_dec(x_313); +lean_dec(x_312); +lean_dec(x_311); +lean_dec(x_310); +lean_dec(x_309); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_320 = lean_ctor_get(x_313, 0); -lean_inc(x_320); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - x_321 = x_313; +x_379 = lean_ctor_get(x_372, 0); +lean_inc(x_379); +if (lean_is_exclusive(x_372)) { + lean_ctor_release(x_372, 0); + x_380 = x_372; } else { - lean_dec_ref(x_313); - x_321 = lean_box(0); + lean_dec_ref(x_372); + x_380 = lean_box(0); } -if (lean_is_scalar(x_321)) { - x_322 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_380)) { + x_381 = lean_alloc_ctor(1, 1, 0); } else { - x_322 = x_321; + x_381 = x_380; } -lean_ctor_set(x_322, 0, x_320); -return x_322; +lean_ctor_set(x_381, 0, x_379); +return x_381; } } case 11: { -lean_object* x_323; lean_object* x_324; -lean_dec(x_256); -lean_dec_ref(x_249); -x_323 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_324 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_323, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_324) == 0) -{ -lean_object* x_325; uint8_t x_326; uint8_t x_327; uint8_t x_328; uint8_t x_329; lean_object* x_330; -x_325 = lean_ctor_get(x_324, 0); -lean_inc(x_325); -lean_dec_ref(x_324); -x_326 = lean_unbox(x_251); -lean_dec(x_251); -x_327 = lean_unbox(x_252); -lean_dec(x_252); -x_328 = lean_unbox(x_253); -lean_dec(x_253); -x_329 = lean_unbox(x_254); -lean_dec(x_254); -x_330 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_326, x_327, x_328, x_329, x_255, x_325, x_5, x_6, x_7, x_8); -x_260 = x_330; -goto block_264; -} -else -{ -lean_object* x_331; lean_object* x_332; lean_object* x_333; -lean_dec_ref(x_3); -lean_dec(x_255); -lean_dec(x_254); -lean_dec(x_253); -lean_dec(x_252); -lean_dec(x_251); -lean_dec(x_250); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_382; lean_object* x_383; +lean_dec(x_315); +lean_dec_ref(x_308); +x_382 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_383 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_382, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_383) == 0) +{ +lean_object* x_384; uint8_t x_385; uint8_t x_386; uint8_t x_387; uint8_t x_388; lean_object* x_389; +x_384 = lean_ctor_get(x_383, 0); +lean_inc(x_384); +lean_dec_ref(x_383); +x_385 = lean_unbox(x_310); +lean_dec(x_310); +x_386 = lean_unbox(x_311); +lean_dec(x_311); +x_387 = lean_unbox(x_312); +lean_dec(x_312); +x_388 = lean_unbox(x_313); +lean_dec(x_313); +x_389 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_385, x_386, x_387, x_388, x_314, x_384, x_6, x_7, x_8, x_9); +x_319 = x_389; +goto block_323; +} +else +{ +lean_object* x_390; lean_object* x_391; lean_object* x_392; +lean_dec_ref(x_4); +lean_dec(x_314); +lean_dec(x_313); +lean_dec(x_312); +lean_dec(x_311); +lean_dec(x_310); +lean_dec(x_309); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_331 = lean_ctor_get(x_324, 0); -lean_inc(x_331); -if (lean_is_exclusive(x_324)) { - lean_ctor_release(x_324, 0); - x_332 = x_324; +x_390 = lean_ctor_get(x_383, 0); +lean_inc(x_390); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + x_391 = x_383; } else { - lean_dec_ref(x_324); - x_332 = lean_box(0); + lean_dec_ref(x_383); + x_391 = lean_box(0); } -if (lean_is_scalar(x_332)) { - x_333 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_391)) { + x_392 = lean_alloc_ctor(1, 1, 0); } else { - x_333 = x_332; + x_392 = x_391; } -lean_ctor_set(x_333, 0, x_331); -return x_333; +lean_ctor_set(x_392, 0, x_390); +return x_392; } } case 13: { -lean_object* x_334; uint8_t x_335; uint8_t x_336; uint8_t x_337; uint8_t x_338; lean_object* x_339; -lean_dec(x_256); -x_334 = lean_box(4); -x_335 = lean_unbox(x_251); -lean_dec(x_251); -x_336 = lean_unbox(x_252); -lean_dec(x_252); -x_337 = lean_unbox(x_253); -lean_dec(x_253); -x_338 = lean_unbox(x_254); -lean_dec(x_254); -x_339 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_335, x_336, x_337, x_338, x_255, x_334, x_5, x_6, x_7, x_8); -x_260 = x_339; -goto block_264; +lean_object* x_393; uint8_t x_394; uint8_t x_395; uint8_t x_396; uint8_t x_397; lean_object* x_398; +lean_dec(x_315); +x_393 = lean_box(4); +x_394 = lean_unbox(x_310); +lean_dec(x_310); +x_395 = lean_unbox(x_311); +lean_dec(x_311); +x_396 = lean_unbox(x_312); +lean_dec(x_312); +x_397 = lean_unbox(x_313); +lean_dec(x_313); +x_398 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_394, x_395, x_396, x_397, x_314, x_393, x_6, x_7, x_8, x_9); +x_319 = x_398; +goto block_323; +} +case 14: +{ +lean_object* x_399; uint8_t x_400; uint8_t x_401; uint8_t x_402; lean_object* x_403; +lean_dec(x_315); +lean_dec(x_310); +lean_inc(x_2); +x_399 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_399, 0, x_316); +lean_ctor_set(x_399, 1, x_2); +lean_ctor_set(x_399, 2, x_308); +x_400 = lean_unbox(x_311); +lean_dec(x_311); +x_401 = lean_unbox(x_312); +lean_dec(x_312); +x_402 = lean_unbox(x_313); +lean_dec(x_313); +x_403 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_292, x_400, x_401, x_402, x_314, x_399, x_6, x_7, x_8, x_9); +x_319 = x_403; +goto block_323; +} +case 15: +{ +lean_object* x_404; lean_object* x_405; uint8_t x_406; uint8_t x_407; uint8_t x_408; lean_object* x_409; +lean_dec(x_315); +lean_dec(x_311); +x_404 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_405 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_405, 0, x_404); +lean_ctor_set(x_405, 1, x_2); +lean_ctor_set(x_405, 2, x_308); +x_406 = lean_unbox(x_310); +lean_dec(x_310); +x_407 = lean_unbox(x_312); +lean_dec(x_312); +x_408 = lean_unbox(x_313); +lean_dec(x_313); +x_409 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_406, x_292, x_407, x_408, x_314, x_405, x_6, x_7, x_8, x_9); +x_319 = x_409; +goto block_323; +} +case 16: +{ +lean_object* x_410; lean_object* x_411; uint8_t x_412; uint8_t x_413; uint8_t x_414; lean_object* x_415; +lean_dec(x_315); +lean_dec(x_312); +x_410 = lean_unsigned_to_nat(4u); +lean_inc(x_2); +x_411 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_411, 0, x_410); +lean_ctor_set(x_411, 1, x_2); +lean_ctor_set(x_411, 2, x_308); +x_412 = lean_unbox(x_310); +lean_dec(x_310); +x_413 = lean_unbox(x_311); +lean_dec(x_311); +x_414 = lean_unbox(x_313); +lean_dec(x_313); +x_415 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_412, x_413, x_292, x_414, x_314, x_411, x_6, x_7, x_8, x_9); +x_319 = x_415; +goto block_323; +} +case 17: +{ +lean_object* x_416; lean_object* x_417; uint8_t x_418; uint8_t x_419; uint8_t x_420; lean_object* x_421; +lean_dec(x_315); +lean_dec(x_313); +x_416 = lean_unsigned_to_nat(8u); +lean_inc(x_2); +x_417 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_417, 0, x_416); +lean_ctor_set(x_417, 1, x_2); +lean_ctor_set(x_417, 2, x_308); +x_418 = lean_unbox(x_310); +lean_dec(x_310); +x_419 = lean_unbox(x_311); +lean_dec(x_311); +x_420 = lean_unbox(x_312); +lean_dec(x_312); +x_421 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_418, x_419, x_420, x_292, x_314, x_417, x_6, x_7, x_8, x_9); +x_319 = x_421; +goto block_323; +} +case 18: +{ +lean_object* x_422; uint8_t x_423; uint8_t x_424; uint8_t x_425; uint8_t x_426; lean_object* x_427; +lean_dec(x_315); +lean_inc(x_2); +x_422 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_422, 0, x_2); +lean_ctor_set_uint8(x_422, sizeof(void*)*1, x_292); +x_423 = lean_unbox(x_310); +lean_dec(x_310); +x_424 = lean_unbox(x_311); +lean_dec(x_311); +x_425 = lean_unbox(x_312); +lean_dec(x_312); +x_426 = lean_unbox(x_313); +lean_dec(x_313); +x_427 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_423, x_424, x_425, x_426, x_314, x_422, x_6, x_7, x_8, x_9); +x_319 = x_427; +goto block_323; } default: { -lean_object* x_340; lean_object* x_341; uint8_t x_342; uint8_t x_343; uint8_t x_344; uint8_t x_345; lean_object* x_346; -x_340 = lean_nat_add(x_255, x_257); -if (lean_is_scalar(x_256)) { - x_341 = lean_alloc_ctor(1, 2, 0); +lean_object* x_428; lean_object* x_429; uint8_t x_430; uint8_t x_431; uint8_t x_432; uint8_t x_433; lean_object* x_434; +x_428 = lean_nat_add(x_314, x_316); +if (lean_is_scalar(x_315)) { + x_429 = lean_alloc_ctor(1, 2, 0); } else { - x_341 = x_256; - lean_ctor_set_tag(x_341, 1); -} -lean_ctor_set(x_341, 0, x_255); -lean_ctor_set(x_341, 1, x_249); -x_342 = lean_unbox(x_251); -lean_dec(x_251); -x_343 = lean_unbox(x_252); -lean_dec(x_252); -x_344 = lean_unbox(x_253); -lean_dec(x_253); -x_345 = lean_unbox(x_254); -lean_dec(x_254); -x_346 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_250, x_342, x_343, x_344, x_345, x_340, x_341, x_5, x_6, x_7, x_8); -x_260 = x_346; -goto block_264; -} -} -block_264: -{ -lean_object* x_261; lean_object* x_262; -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -lean_dec_ref(x_260); -x_262 = lean_ctor_get(x_261, 0); -lean_inc(x_262); -lean_dec(x_261); -x_4 = x_262; + x_429 = x_315; + lean_ctor_set_tag(x_429, 1); +} +lean_ctor_set(x_429, 0, x_314); +lean_ctor_set(x_429, 1, x_308); +x_430 = lean_unbox(x_310); +lean_dec(x_310); +x_431 = lean_unbox(x_311); +lean_dec(x_311); +x_432 = lean_unbox(x_312); +lean_dec(x_312); +x_433 = lean_unbox(x_313); +lean_dec(x_313); +x_434 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_309, x_430, x_431, x_432, x_433, x_428, x_429, x_6, x_7, x_8, x_9); +x_319 = x_434; +goto block_323; +} +} +block_323: +{ +lean_object* x_320; lean_object* x_321; +x_320 = lean_ctor_get(x_319, 0); +lean_inc(x_320); +lean_dec_ref(x_319); +x_321 = lean_ctor_get(x_320, 0); +lean_inc(x_321); +lean_dec(x_320); +x_5 = x_321; goto _start; } } else { -lean_object* x_347; lean_object* x_348; lean_object* x_349; -lean_dec(x_232); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_435; lean_object* x_436; lean_object* x_437; +lean_dec(x_291); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_347 = lean_ctor_get(x_244, 0); -lean_inc(x_347); -if (lean_is_exclusive(x_244)) { - lean_ctor_release(x_244, 0); - x_348 = x_244; +x_435 = lean_ctor_get(x_303, 0); +lean_inc(x_435); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + x_436 = x_303; } else { - lean_dec_ref(x_244); - x_348 = lean_box(0); + lean_dec_ref(x_303); + x_436 = lean_box(0); } -if (lean_is_scalar(x_348)) { - x_349 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_436)) { + x_437 = lean_alloc_ctor(1, 1, 0); } else { - x_349 = x_348; + x_437 = x_436; } -lean_ctor_set(x_349, 0, x_347); -return x_349; +lean_ctor_set(x_437, 0, x_435); +return x_437; } } else { -lean_object* x_350; lean_object* x_351; lean_object* x_352; -lean_dec(x_232); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_438; lean_object* x_439; lean_object* x_440; +lean_dec(x_291); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_350 = lean_ctor_get(x_242, 0); -lean_inc(x_350); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - x_351 = x_242; +x_438 = lean_ctor_get(x_301, 0); +lean_inc(x_438); +if (lean_is_exclusive(x_301)) { + lean_ctor_release(x_301, 0); + x_439 = x_301; } else { - lean_dec_ref(x_242); - x_351 = lean_box(0); + lean_dec_ref(x_301); + x_439 = lean_box(0); } -if (lean_is_scalar(x_351)) { - x_352 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_439)) { + x_440 = lean_alloc_ctor(1, 1, 0); } else { - x_352 = x_351; + x_440 = x_439; } -lean_ctor_set(x_352, 0, x_350); -return x_352; +lean_ctor_set(x_440, 0, x_438); +return x_440; } } else { -lean_object* x_353; lean_object* x_354; lean_object* x_355; -lean_dec(x_232); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_441; lean_object* x_442; lean_object* x_443; +lean_dec(x_291); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_353 = lean_ctor_get(x_240, 0); -lean_inc(x_353); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - x_354 = x_240; +x_441 = lean_ctor_get(x_299, 0); +lean_inc(x_441); +if (lean_is_exclusive(x_299)) { + lean_ctor_release(x_299, 0); + x_442 = x_299; } else { - lean_dec_ref(x_240); - x_354 = lean_box(0); + lean_dec_ref(x_299); + x_442 = lean_box(0); } -if (lean_is_scalar(x_354)) { - x_355 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_442)) { + x_443 = lean_alloc_ctor(1, 1, 0); } else { - x_355 = x_354; + x_443 = x_442; } -lean_ctor_set(x_355, 0, x_353); -return x_355; +lean_ctor_set(x_443, 0, x_441); +return x_443; } } else { -lean_object* x_356; lean_object* x_357; lean_object* x_358; -lean_dec(x_232); -lean_free_object(x_3); -lean_dec(x_13); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_444; lean_object* x_445; lean_object* x_446; +lean_dec(x_291); +lean_free_object(x_4); +lean_dec(x_14); +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_2); -x_356 = lean_ctor_get(x_238, 0); -lean_inc(x_356); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - x_357 = x_238; +x_444 = lean_ctor_get(x_297, 0); +lean_inc(x_444); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + x_445 = x_297; } else { - lean_dec_ref(x_238); - x_357 = lean_box(0); + lean_dec_ref(x_297); + x_445 = lean_box(0); } -if (lean_is_scalar(x_357)) { - x_358 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_445)) { + x_446 = lean_alloc_ctor(1, 1, 0); } else { - x_358 = x_357; + x_446 = x_445; } -lean_ctor_set(x_358, 0, x_356); -return x_358; +lean_ctor_set(x_446, 0, x_444); +return x_446; } } } } else { -lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; -x_359 = lean_ctor_get(x_3, 1); -lean_inc(x_359); -lean_dec(x_3); -x_360 = lean_ctor_get(x_10, 0); -lean_inc(x_360); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - x_361 = x_10; +lean_object* x_447; lean_object* x_448; lean_object* x_449; uint8_t x_450; +x_447 = lean_ctor_get(x_4, 1); +lean_inc(x_447); +lean_dec(x_4); +x_448 = lean_ctor_get(x_11, 0); +lean_inc(x_448); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + x_449 = x_11; } else { - lean_dec_ref(x_10); - x_361 = lean_box(0); + lean_dec_ref(x_11); + x_449 = lean_box(0); } -x_362 = lean_nat_dec_lt(x_360, x_359); -if (x_362 == 0) +x_450 = lean_nat_dec_lt(x_448, x_447); +if (x_450 == 0) { -lean_object* x_363; -lean_dec(x_361); -lean_dec(x_360); -lean_dec(x_359); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_451; +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_447); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_363 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_363, 0, x_4); -return x_363; +x_451 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_451, 0, x_5); +return x_451; } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -x_364 = lean_ctor_get(x_1, 0); -x_365 = lean_array_fget_borrowed(x_364, x_360); -x_366 = l_Lean_Expr_fvarId_x21(x_365); -lean_inc_ref(x_5); -x_367 = l_Lean_FVarId_getType___redArg(x_366, x_5, x_7, x_8); -if (lean_obj_tag(x_367) == 0) +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; +x_452 = lean_ctor_get(x_1, 0); +x_453 = lean_array_fget_borrowed(x_452, x_448); +x_454 = l_Lean_Expr_fvarId_x21(x_453); +lean_inc_ref(x_6); +x_455 = l_Lean_FVarId_getType___redArg(x_454, x_6, x_8, x_9); +if (lean_obj_tag(x_455) == 0) { -lean_object* x_368; lean_object* x_369; -x_368 = lean_ctor_get(x_367, 0); -lean_inc(x_368); -lean_dec_ref(x_367); -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_369 = l_Lean_Compiler_LCNF_toLCNFType(x_368, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_369) == 0) -{ -lean_object* x_370; lean_object* x_371; -x_370 = lean_ctor_get(x_369, 0); -lean_inc(x_370); -lean_dec_ref(x_369); -lean_inc(x_8); -lean_inc_ref(x_7); -x_371 = l_Lean_Compiler_LCNF_toMonoType(x_370, x_7, x_8); -if (lean_obj_tag(x_371) == 0) -{ -lean_object* x_372; lean_object* x_373; -x_372 = lean_ctor_get(x_371, 0); -lean_inc(x_372); -lean_dec_ref(x_371); -lean_inc(x_8); -lean_inc_ref(x_7); -x_373 = l_Lean_IR_toIRType(x_372, x_7, x_8); -if (lean_obj_tag(x_373) == 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_object* x_390; -x_374 = lean_ctor_get(x_4, 1); -lean_inc(x_374); -x_375 = lean_ctor_get(x_374, 1); -lean_inc(x_375); -x_376 = lean_ctor_get(x_375, 1); -lean_inc(x_376); -x_377 = lean_ctor_get(x_376, 1); -lean_inc(x_377); -x_378 = lean_ctor_get(x_373, 0); -lean_inc(x_378); -lean_dec_ref(x_373); -x_379 = lean_ctor_get(x_4, 0); -lean_inc(x_379); -lean_dec_ref(x_4); -x_380 = lean_ctor_get(x_374, 0); -lean_inc(x_380); -lean_dec(x_374); -x_381 = lean_ctor_get(x_375, 0); -lean_inc(x_381); -lean_dec(x_375); -x_382 = lean_ctor_get(x_376, 0); -lean_inc(x_382); -lean_dec(x_376); -x_383 = lean_ctor_get(x_377, 0); -lean_inc(x_383); -x_384 = lean_ctor_get(x_377, 1); -lean_inc(x_384); -if (lean_is_exclusive(x_377)) { - lean_ctor_release(x_377, 0); - lean_ctor_release(x_377, 1); - x_385 = x_377; +lean_object* x_456; lean_object* x_457; +x_456 = lean_ctor_get(x_455, 0); +lean_inc(x_456); +lean_dec_ref(x_455); +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_457 = l_Lean_Compiler_LCNF_toLCNFType(x_456, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_457) == 0) +{ +lean_object* x_458; lean_object* x_459; +x_458 = lean_ctor_get(x_457, 0); +lean_inc(x_458); +lean_dec_ref(x_457); +lean_inc(x_9); +lean_inc_ref(x_8); +x_459 = l_Lean_Compiler_LCNF_toMonoType(x_458, x_8, x_9); +if (lean_obj_tag(x_459) == 0) +{ +lean_object* x_460; lean_object* x_461; +x_460 = lean_ctor_get(x_459, 0); +lean_inc(x_460); +lean_dec_ref(x_459); +lean_inc(x_9); +lean_inc_ref(x_8); +x_461 = l_Lean_IR_toIRType(x_460, x_8, x_9); +if (lean_obj_tag(x_461) == 0) +{ +lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; +x_462 = lean_ctor_get(x_5, 1); +lean_inc(x_462); +x_463 = lean_ctor_get(x_462, 1); +lean_inc(x_463); +x_464 = lean_ctor_get(x_463, 1); +lean_inc(x_464); +x_465 = lean_ctor_get(x_464, 1); +lean_inc(x_465); +x_466 = lean_ctor_get(x_461, 0); +lean_inc(x_466); +lean_dec_ref(x_461); +x_467 = lean_ctor_get(x_5, 0); +lean_inc(x_467); +lean_dec_ref(x_5); +x_468 = lean_ctor_get(x_462, 0); +lean_inc(x_468); +lean_dec(x_462); +x_469 = lean_ctor_get(x_463, 0); +lean_inc(x_469); +lean_dec(x_463); +x_470 = lean_ctor_get(x_464, 0); +lean_inc(x_470); +lean_dec(x_464); +x_471 = lean_ctor_get(x_465, 0); +lean_inc(x_471); +x_472 = lean_ctor_get(x_465, 1); +lean_inc(x_472); +if (lean_is_exclusive(x_465)) { + lean_ctor_release(x_465, 0); + lean_ctor_release(x_465, 1); + x_473 = x_465; } else { - lean_dec_ref(x_377); - x_385 = lean_box(0); -} -x_386 = lean_unsigned_to_nat(1u); -x_387 = lean_nat_add(x_360, x_386); -lean_dec(x_360); -if (lean_is_scalar(x_361)) { - x_388 = lean_alloc_ctor(1, 1, 0); + lean_dec_ref(x_465); + x_473 = lean_box(0); +} +x_474 = lean_unsigned_to_nat(1u); +x_475 = lean_nat_add(x_448, x_474); +lean_dec(x_448); +if (lean_is_scalar(x_449)) { + x_476 = lean_alloc_ctor(1, 1, 0); } else { - x_388 = x_361; + x_476 = x_449; } -lean_ctor_set(x_388, 0, x_387); -x_389 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_389, 0, x_388); -lean_ctor_set(x_389, 1, x_359); -switch (lean_obj_tag(x_378)) { +lean_ctor_set(x_476, 0, x_475); +x_477 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_477, 0, x_476); +lean_ctor_set(x_477, 1, x_447); +switch (lean_obj_tag(x_466)) { case 0: { -lean_object* x_395; lean_object* x_396; uint8_t x_397; uint8_t x_398; uint8_t x_399; lean_object* x_400; -lean_dec(x_385); -lean_dec(x_383); -x_395 = lean_unsigned_to_nat(8u); +lean_object* x_483; lean_object* x_484; uint8_t x_485; uint8_t x_486; uint8_t x_487; lean_object* x_488; +lean_dec(x_473); +lean_dec(x_471); +x_483 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_396 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_396, 0, x_395); -lean_ctor_set(x_396, 1, x_2); -lean_ctor_set(x_396, 2, x_378); -x_397 = lean_unbox(x_380); -lean_dec(x_380); -x_398 = lean_unbox(x_381); -lean_dec(x_381); -x_399 = lean_unbox(x_382); -lean_dec(x_382); -x_400 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_397, x_398, x_399, x_362, x_384, x_396, x_5, x_6, x_7, x_8); -x_390 = x_400; -goto block_394; +x_484 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_484, 0, x_483); +lean_ctor_set(x_484, 1, x_2); +lean_ctor_set(x_484, 2, x_466); +x_485 = lean_unbox(x_468); +lean_dec(x_468); +x_486 = lean_unbox(x_469); +lean_dec(x_469); +x_487 = lean_unbox(x_470); +lean_dec(x_470); +x_488 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_485, x_486, x_487, x_450, x_472, x_484, x_6, x_7, x_8, x_9); +x_478 = x_488; +goto block_482; } case 1: { -lean_object* x_401; uint8_t x_402; uint8_t x_403; uint8_t x_404; lean_object* x_405; -lean_dec(x_385); -lean_dec(x_380); +lean_object* x_489; uint8_t x_490; uint8_t x_491; uint8_t x_492; lean_object* x_493; +lean_dec(x_473); +lean_dec(x_468); lean_inc(x_2); -x_401 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_401, 0, x_386); -lean_ctor_set(x_401, 1, x_2); -lean_ctor_set(x_401, 2, x_378); -x_402 = lean_unbox(x_381); -lean_dec(x_381); -x_403 = lean_unbox(x_382); -lean_dec(x_382); -x_404 = lean_unbox(x_383); -lean_dec(x_383); -x_405 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_362, x_402, x_403, x_404, x_384, x_401, x_5, x_6, x_7, x_8); -x_390 = x_405; -goto block_394; +x_489 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_489, 0, x_474); +lean_ctor_set(x_489, 1, x_2); +lean_ctor_set(x_489, 2, x_466); +x_490 = lean_unbox(x_469); +lean_dec(x_469); +x_491 = lean_unbox(x_470); +lean_dec(x_470); +x_492 = lean_unbox(x_471); +lean_dec(x_471); +x_493 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_450, x_490, x_491, x_492, x_472, x_489, x_6, x_7, x_8, x_9); +x_478 = x_493; +goto block_482; } case 2: { -lean_object* x_406; lean_object* x_407; uint8_t x_408; uint8_t x_409; uint8_t x_410; lean_object* x_411; -lean_dec(x_385); -lean_dec(x_381); -x_406 = lean_unsigned_to_nat(2u); +lean_object* x_494; lean_object* x_495; uint8_t x_496; uint8_t x_497; uint8_t x_498; lean_object* x_499; +lean_dec(x_473); +lean_dec(x_469); +x_494 = lean_unsigned_to_nat(2u); lean_inc(x_2); -x_407 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_407, 0, x_406); -lean_ctor_set(x_407, 1, x_2); -lean_ctor_set(x_407, 2, x_378); -x_408 = lean_unbox(x_380); -lean_dec(x_380); -x_409 = lean_unbox(x_382); -lean_dec(x_382); -x_410 = lean_unbox(x_383); -lean_dec(x_383); -x_411 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_408, x_362, x_409, x_410, x_384, x_407, x_5, x_6, x_7, x_8); -x_390 = x_411; -goto block_394; +x_495 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_495, 0, x_494); +lean_ctor_set(x_495, 1, x_2); +lean_ctor_set(x_495, 2, x_466); +x_496 = lean_unbox(x_468); +lean_dec(x_468); +x_497 = lean_unbox(x_470); +lean_dec(x_470); +x_498 = lean_unbox(x_471); +lean_dec(x_471); +x_499 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_496, x_450, x_497, x_498, x_472, x_495, x_6, x_7, x_8, x_9); +x_478 = x_499; +goto block_482; } case 3: { -lean_object* x_412; lean_object* x_413; uint8_t x_414; uint8_t x_415; uint8_t x_416; lean_object* x_417; -lean_dec(x_385); -lean_dec(x_382); -x_412 = lean_unsigned_to_nat(4u); +lean_object* x_500; lean_object* x_501; uint8_t x_502; uint8_t x_503; uint8_t x_504; lean_object* x_505; +lean_dec(x_473); +lean_dec(x_470); +x_500 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_413 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_413, 0, x_412); -lean_ctor_set(x_413, 1, x_2); -lean_ctor_set(x_413, 2, x_378); -x_414 = lean_unbox(x_380); -lean_dec(x_380); -x_415 = lean_unbox(x_381); -lean_dec(x_381); -x_416 = lean_unbox(x_383); -lean_dec(x_383); -x_417 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_414, x_415, x_362, x_416, x_384, x_413, x_5, x_6, x_7, x_8); -x_390 = x_417; -goto block_394; +x_501 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_501, 0, x_500); +lean_ctor_set(x_501, 1, x_2); +lean_ctor_set(x_501, 2, x_466); +x_502 = lean_unbox(x_468); +lean_dec(x_468); +x_503 = lean_unbox(x_469); +lean_dec(x_469); +x_504 = lean_unbox(x_471); +lean_dec(x_471); +x_505 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_502, x_503, x_450, x_504, x_472, x_501, x_6, x_7, x_8, x_9); +x_478 = x_505; +goto block_482; } case 4: { -lean_object* x_418; lean_object* x_419; uint8_t x_420; uint8_t x_421; uint8_t x_422; lean_object* x_423; -lean_dec(x_385); -lean_dec(x_383); -x_418 = lean_unsigned_to_nat(8u); +lean_object* x_506; lean_object* x_507; uint8_t x_508; uint8_t x_509; uint8_t x_510; lean_object* x_511; +lean_dec(x_473); +lean_dec(x_471); +x_506 = lean_unsigned_to_nat(8u); lean_inc(x_2); -x_419 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_419, 0, x_418); -lean_ctor_set(x_419, 1, x_2); -lean_ctor_set(x_419, 2, x_378); -x_420 = lean_unbox(x_380); -lean_dec(x_380); -x_421 = lean_unbox(x_381); -lean_dec(x_381); -x_422 = lean_unbox(x_382); -lean_dec(x_382); -x_423 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_420, x_421, x_422, x_362, x_384, x_419, x_5, x_6, x_7, x_8); -x_390 = x_423; -goto block_394; +x_507 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_507, 0, x_506); +lean_ctor_set(x_507, 1, x_2); +lean_ctor_set(x_507, 2, x_466); +x_508 = lean_unbox(x_468); +lean_dec(x_468); +x_509 = lean_unbox(x_469); +lean_dec(x_469); +x_510 = lean_unbox(x_470); +lean_dec(x_470); +x_511 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_508, x_509, x_510, x_450, x_472, x_507, x_6, x_7, x_8, x_9); +x_478 = x_511; +goto block_482; } case 5: { -lean_object* x_424; uint8_t x_425; uint8_t x_426; uint8_t x_427; uint8_t x_428; lean_object* x_429; -lean_dec(x_385); +lean_object* x_512; uint8_t x_513; uint8_t x_514; uint8_t x_515; uint8_t x_516; lean_object* x_517; +lean_dec(x_473); lean_inc(x_2); -x_424 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_424, 0, x_2); -x_425 = lean_unbox(x_380); -lean_dec(x_380); -x_426 = lean_unbox(x_381); -lean_dec(x_381); -x_427 = lean_unbox(x_382); -lean_dec(x_382); -x_428 = lean_unbox(x_383); -lean_dec(x_383); -x_429 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_425, x_426, x_427, x_428, x_384, x_424, x_5, x_6, x_7, x_8); -x_390 = x_429; -goto block_394; +x_512 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_512, 0, x_2); +lean_ctor_set_uint8(x_512, sizeof(void*)*1, x_3); +x_513 = lean_unbox(x_468); +lean_dec(x_468); +x_514 = lean_unbox(x_469); +lean_dec(x_469); +x_515 = lean_unbox(x_470); +lean_dec(x_470); +x_516 = lean_unbox(x_471); +lean_dec(x_471); +x_517 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_513, x_514, x_515, x_516, x_472, x_512, x_6, x_7, x_8, x_9); +x_478 = x_517; +goto block_482; } case 6: { -lean_object* x_430; uint8_t x_431; uint8_t x_432; uint8_t x_433; uint8_t x_434; lean_object* x_435; -lean_dec(x_385); -x_430 = lean_box(0); -x_431 = lean_unbox(x_380); -lean_dec(x_380); -x_432 = lean_unbox(x_381); -lean_dec(x_381); -x_433 = lean_unbox(x_382); -lean_dec(x_382); -x_434 = lean_unbox(x_383); -lean_dec(x_383); -x_435 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_431, x_432, x_433, x_434, x_384, x_430, x_5, x_6, x_7, x_8); -x_390 = x_435; -goto block_394; +lean_object* x_518; uint8_t x_519; uint8_t x_520; uint8_t x_521; uint8_t x_522; lean_object* x_523; +lean_dec(x_473); +x_518 = lean_box(0); +x_519 = lean_unbox(x_468); +lean_dec(x_468); +x_520 = lean_unbox(x_469); +lean_dec(x_469); +x_521 = lean_unbox(x_470); +lean_dec(x_470); +x_522 = lean_unbox(x_471); +lean_dec(x_471); +x_523 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_519, x_520, x_521, x_522, x_472, x_518, x_6, x_7, x_8, x_9); +x_478 = x_523; +goto block_482; } case 9: { -lean_object* x_436; lean_object* x_437; uint8_t x_438; uint8_t x_439; uint8_t x_440; lean_object* x_441; -lean_dec(x_385); -lean_dec(x_382); -x_436 = lean_unsigned_to_nat(4u); +lean_object* x_524; lean_object* x_525; uint8_t x_526; uint8_t x_527; uint8_t x_528; lean_object* x_529; +lean_dec(x_473); +lean_dec(x_470); +x_524 = lean_unsigned_to_nat(4u); lean_inc(x_2); -x_437 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_437, 0, x_436); -lean_ctor_set(x_437, 1, x_2); -lean_ctor_set(x_437, 2, x_378); -x_438 = lean_unbox(x_380); -lean_dec(x_380); -x_439 = lean_unbox(x_381); -lean_dec(x_381); -x_440 = lean_unbox(x_383); -lean_dec(x_383); -x_441 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_438, x_439, x_362, x_440, x_384, x_437, x_5, x_6, x_7, x_8); -x_390 = x_441; -goto block_394; +x_525 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_525, 0, x_524); +lean_ctor_set(x_525, 1, x_2); +lean_ctor_set(x_525, 2, x_466); +x_526 = lean_unbox(x_468); +lean_dec(x_468); +x_527 = lean_unbox(x_469); +lean_dec(x_469); +x_528 = lean_unbox(x_471); +lean_dec(x_471); +x_529 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_526, x_527, x_450, x_528, x_472, x_525, x_6, x_7, x_8, x_9); +x_478 = x_529; +goto block_482; } case 10: { -lean_object* x_442; lean_object* x_443; -lean_dec(x_385); -lean_dec_ref(x_378); -x_442 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_443 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_442, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_443) == 0) -{ -lean_object* x_444; uint8_t x_445; uint8_t x_446; uint8_t x_447; uint8_t x_448; lean_object* x_449; -x_444 = lean_ctor_get(x_443, 0); -lean_inc(x_444); -lean_dec_ref(x_443); -x_445 = lean_unbox(x_380); -lean_dec(x_380); -x_446 = lean_unbox(x_381); -lean_dec(x_381); -x_447 = lean_unbox(x_382); -lean_dec(x_382); -x_448 = lean_unbox(x_383); -lean_dec(x_383); -x_449 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_445, x_446, x_447, x_448, x_384, x_444, x_5, x_6, x_7, x_8); -x_390 = x_449; -goto block_394; -} -else -{ -lean_object* x_450; lean_object* x_451; lean_object* x_452; -lean_dec_ref(x_389); -lean_dec(x_384); -lean_dec(x_383); -lean_dec(x_382); -lean_dec(x_381); -lean_dec(x_380); -lean_dec(x_379); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_530; lean_object* x_531; +lean_dec(x_473); +lean_dec_ref(x_466); +x_530 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_531 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_530, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_531) == 0) +{ +lean_object* x_532; uint8_t x_533; uint8_t x_534; uint8_t x_535; uint8_t x_536; lean_object* x_537; +x_532 = lean_ctor_get(x_531, 0); +lean_inc(x_532); +lean_dec_ref(x_531); +x_533 = lean_unbox(x_468); +lean_dec(x_468); +x_534 = lean_unbox(x_469); +lean_dec(x_469); +x_535 = lean_unbox(x_470); +lean_dec(x_470); +x_536 = lean_unbox(x_471); +lean_dec(x_471); +x_537 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_533, x_534, x_535, x_536, x_472, x_532, x_6, x_7, x_8, x_9); +x_478 = x_537; +goto block_482; +} +else +{ +lean_object* x_538; lean_object* x_539; lean_object* x_540; +lean_dec_ref(x_477); +lean_dec(x_472); +lean_dec(x_471); +lean_dec(x_470); +lean_dec(x_469); +lean_dec(x_468); +lean_dec(x_467); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_450 = lean_ctor_get(x_443, 0); -lean_inc(x_450); -if (lean_is_exclusive(x_443)) { - lean_ctor_release(x_443, 0); - x_451 = x_443; +x_538 = lean_ctor_get(x_531, 0); +lean_inc(x_538); +if (lean_is_exclusive(x_531)) { + lean_ctor_release(x_531, 0); + x_539 = x_531; } else { - lean_dec_ref(x_443); - x_451 = lean_box(0); + lean_dec_ref(x_531); + x_539 = lean_box(0); } -if (lean_is_scalar(x_451)) { - x_452 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_539)) { + x_540 = lean_alloc_ctor(1, 1, 0); } else { - x_452 = x_451; + x_540 = x_539; } -lean_ctor_set(x_452, 0, x_450); -return x_452; +lean_ctor_set(x_540, 0, x_538); +return x_540; } } case 11: { -lean_object* x_453; lean_object* x_454; -lean_dec(x_385); -lean_dec_ref(x_378); -x_453 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; -lean_inc(x_8); -lean_inc_ref(x_7); -lean_inc(x_6); -lean_inc_ref(x_5); -x_454 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_453, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_454) == 0) -{ -lean_object* x_455; uint8_t x_456; uint8_t x_457; uint8_t x_458; uint8_t x_459; lean_object* x_460; -x_455 = lean_ctor_get(x_454, 0); -lean_inc(x_455); -lean_dec_ref(x_454); -x_456 = lean_unbox(x_380); -lean_dec(x_380); -x_457 = lean_unbox(x_381); -lean_dec(x_381); -x_458 = lean_unbox(x_382); -lean_dec(x_382); -x_459 = lean_unbox(x_383); -lean_dec(x_383); -x_460 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_456, x_457, x_458, x_459, x_384, x_455, x_5, x_6, x_7, x_8); -x_390 = x_460; -goto block_394; -} -else -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; -lean_dec_ref(x_389); -lean_dec(x_384); -lean_dec(x_383); -lean_dec(x_382); -lean_dec(x_381); -lean_dec(x_380); -lean_dec(x_379); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); -lean_dec_ref(x_5); +lean_object* x_541; lean_object* x_542; +lean_dec(x_473); +lean_dec_ref(x_466); +x_541 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__1; +lean_inc(x_9); +lean_inc_ref(x_8); +lean_inc(x_7); +lean_inc_ref(x_6); +x_542 = l_panic___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__2(x_541, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_542) == 0) +{ +lean_object* x_543; uint8_t x_544; uint8_t x_545; uint8_t x_546; uint8_t x_547; lean_object* x_548; +x_543 = lean_ctor_get(x_542, 0); +lean_inc(x_543); +lean_dec_ref(x_542); +x_544 = lean_unbox(x_468); +lean_dec(x_468); +x_545 = lean_unbox(x_469); +lean_dec(x_469); +x_546 = lean_unbox(x_470); +lean_dec(x_470); +x_547 = lean_unbox(x_471); +lean_dec(x_471); +x_548 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_544, x_545, x_546, x_547, x_472, x_543, x_6, x_7, x_8, x_9); +x_478 = x_548; +goto block_482; +} +else +{ +lean_object* x_549; lean_object* x_550; lean_object* x_551; +lean_dec_ref(x_477); +lean_dec(x_472); +lean_dec(x_471); +lean_dec(x_470); +lean_dec(x_469); +lean_dec(x_468); +lean_dec(x_467); +lean_dec(x_9); +lean_dec_ref(x_8); +lean_dec(x_7); +lean_dec_ref(x_6); lean_dec(x_2); -x_461 = lean_ctor_get(x_454, 0); -lean_inc(x_461); -if (lean_is_exclusive(x_454)) { - lean_ctor_release(x_454, 0); - x_462 = x_454; +x_549 = lean_ctor_get(x_542, 0); +lean_inc(x_549); +if (lean_is_exclusive(x_542)) { + lean_ctor_release(x_542, 0); + x_550 = x_542; } else { - lean_dec_ref(x_454); - x_462 = lean_box(0); + lean_dec_ref(x_542); + x_550 = lean_box(0); } -if (lean_is_scalar(x_462)) { - x_463 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_550)) { + x_551 = lean_alloc_ctor(1, 1, 0); } else { - x_463 = x_462; + x_551 = x_550; } -lean_ctor_set(x_463, 0, x_461); -return x_463; +lean_ctor_set(x_551, 0, x_549); +return x_551; } } case 13: { -lean_object* x_464; uint8_t x_465; uint8_t x_466; uint8_t x_467; uint8_t x_468; lean_object* x_469; -lean_dec(x_385); -x_464 = lean_box(4); -x_465 = lean_unbox(x_380); -lean_dec(x_380); -x_466 = lean_unbox(x_381); -lean_dec(x_381); -x_467 = lean_unbox(x_382); -lean_dec(x_382); -x_468 = lean_unbox(x_383); -lean_dec(x_383); -x_469 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_465, x_466, x_467, x_468, x_384, x_464, x_5, x_6, x_7, x_8); -x_390 = x_469; -goto block_394; +lean_object* x_552; uint8_t x_553; uint8_t x_554; uint8_t x_555; uint8_t x_556; lean_object* x_557; +lean_dec(x_473); +x_552 = lean_box(4); +x_553 = lean_unbox(x_468); +lean_dec(x_468); +x_554 = lean_unbox(x_469); +lean_dec(x_469); +x_555 = lean_unbox(x_470); +lean_dec(x_470); +x_556 = lean_unbox(x_471); +lean_dec(x_471); +x_557 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_553, x_554, x_555, x_556, x_472, x_552, x_6, x_7, x_8, x_9); +x_478 = x_557; +goto block_482; +} +case 14: +{ +lean_object* x_558; uint8_t x_559; uint8_t x_560; uint8_t x_561; lean_object* x_562; +lean_dec(x_473); +lean_dec(x_468); +lean_inc(x_2); +x_558 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_558, 0, x_474); +lean_ctor_set(x_558, 1, x_2); +lean_ctor_set(x_558, 2, x_466); +x_559 = lean_unbox(x_469); +lean_dec(x_469); +x_560 = lean_unbox(x_470); +lean_dec(x_470); +x_561 = lean_unbox(x_471); +lean_dec(x_471); +x_562 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_450, x_559, x_560, x_561, x_472, x_558, x_6, x_7, x_8, x_9); +x_478 = x_562; +goto block_482; +} +case 15: +{ +lean_object* x_563; lean_object* x_564; uint8_t x_565; uint8_t x_566; uint8_t x_567; lean_object* x_568; +lean_dec(x_473); +lean_dec(x_469); +x_563 = lean_unsigned_to_nat(2u); +lean_inc(x_2); +x_564 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_564, 0, x_563); +lean_ctor_set(x_564, 1, x_2); +lean_ctor_set(x_564, 2, x_466); +x_565 = lean_unbox(x_468); +lean_dec(x_468); +x_566 = lean_unbox(x_470); +lean_dec(x_470); +x_567 = lean_unbox(x_471); +lean_dec(x_471); +x_568 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_565, x_450, x_566, x_567, x_472, x_564, x_6, x_7, x_8, x_9); +x_478 = x_568; +goto block_482; +} +case 16: +{ +lean_object* x_569; lean_object* x_570; uint8_t x_571; uint8_t x_572; uint8_t x_573; lean_object* x_574; +lean_dec(x_473); +lean_dec(x_470); +x_569 = lean_unsigned_to_nat(4u); +lean_inc(x_2); +x_570 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_570, 0, x_569); +lean_ctor_set(x_570, 1, x_2); +lean_ctor_set(x_570, 2, x_466); +x_571 = lean_unbox(x_468); +lean_dec(x_468); +x_572 = lean_unbox(x_469); +lean_dec(x_469); +x_573 = lean_unbox(x_471); +lean_dec(x_471); +x_574 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_571, x_572, x_450, x_573, x_472, x_570, x_6, x_7, x_8, x_9); +x_478 = x_574; +goto block_482; +} +case 17: +{ +lean_object* x_575; lean_object* x_576; uint8_t x_577; uint8_t x_578; uint8_t x_579; lean_object* x_580; +lean_dec(x_473); +lean_dec(x_471); +x_575 = lean_unsigned_to_nat(8u); +lean_inc(x_2); +x_576 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_576, 0, x_575); +lean_ctor_set(x_576, 1, x_2); +lean_ctor_set(x_576, 2, x_466); +x_577 = lean_unbox(x_468); +lean_dec(x_468); +x_578 = lean_unbox(x_469); +lean_dec(x_469); +x_579 = lean_unbox(x_470); +lean_dec(x_470); +x_580 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_577, x_578, x_579, x_450, x_472, x_576, x_6, x_7, x_8, x_9); +x_478 = x_580; +goto block_482; +} +case 18: +{ +lean_object* x_581; uint8_t x_582; uint8_t x_583; uint8_t x_584; uint8_t x_585; lean_object* x_586; +lean_dec(x_473); +lean_inc(x_2); +x_581 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_581, 0, x_2); +lean_ctor_set_uint8(x_581, sizeof(void*)*1, x_450); +x_582 = lean_unbox(x_468); +lean_dec(x_468); +x_583 = lean_unbox(x_469); +lean_dec(x_469); +x_584 = lean_unbox(x_470); +lean_dec(x_470); +x_585 = lean_unbox(x_471); +lean_dec(x_471); +x_586 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_582, x_583, x_584, x_585, x_472, x_581, x_6, x_7, x_8, x_9); +x_478 = x_586; +goto block_482; } default: { -lean_object* x_470; lean_object* x_471; uint8_t x_472; uint8_t x_473; uint8_t x_474; uint8_t x_475; lean_object* x_476; -x_470 = lean_nat_add(x_384, x_386); -if (lean_is_scalar(x_385)) { - x_471 = lean_alloc_ctor(1, 2, 0); +lean_object* x_587; lean_object* x_588; uint8_t x_589; uint8_t x_590; uint8_t x_591; uint8_t x_592; lean_object* x_593; +x_587 = lean_nat_add(x_472, x_474); +if (lean_is_scalar(x_473)) { + x_588 = lean_alloc_ctor(1, 2, 0); } else { - x_471 = x_385; - lean_ctor_set_tag(x_471, 1); -} -lean_ctor_set(x_471, 0, x_384); -lean_ctor_set(x_471, 1, x_378); -x_472 = lean_unbox(x_380); -lean_dec(x_380); -x_473 = lean_unbox(x_381); -lean_dec(x_381); -x_474 = lean_unbox(x_382); -lean_dec(x_382); -x_475 = lean_unbox(x_383); -lean_dec(x_383); -x_476 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_379, x_472, x_473, x_474, x_475, x_470, x_471, x_5, x_6, x_7, x_8); -x_390 = x_476; -goto block_394; -} -} -block_394: -{ -lean_object* x_391; lean_object* x_392; -x_391 = lean_ctor_get(x_390, 0); -lean_inc(x_391); -lean_dec_ref(x_390); -x_392 = lean_ctor_get(x_391, 0); -lean_inc(x_392); -lean_dec(x_391); -x_3 = x_389; -x_4 = x_392; + x_588 = x_473; + lean_ctor_set_tag(x_588, 1); +} +lean_ctor_set(x_588, 0, x_472); +lean_ctor_set(x_588, 1, x_466); +x_589 = lean_unbox(x_468); +lean_dec(x_468); +x_590 = lean_unbox(x_469); +lean_dec(x_469); +x_591 = lean_unbox(x_470); +lean_dec(x_470); +x_592 = lean_unbox(x_471); +lean_dec(x_471); +x_593 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___lam__0(x_467, x_589, x_590, x_591, x_592, x_587, x_588, x_6, x_7, x_8, x_9); +x_478 = x_593; +goto block_482; +} +} +block_482: +{ +lean_object* x_479; lean_object* x_480; +x_479 = lean_ctor_get(x_478, 0); +lean_inc(x_479); +lean_dec_ref(x_478); +x_480 = lean_ctor_get(x_479, 0); +lean_inc(x_480); +lean_dec(x_479); +x_4 = x_477; +x_5 = x_480; goto _start; } } else { -lean_object* x_477; lean_object* x_478; lean_object* x_479; -lean_dec(x_361); -lean_dec(x_360); -lean_dec(x_359); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_594; lean_object* x_595; lean_object* x_596; +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_447); +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_2); -x_477 = lean_ctor_get(x_373, 0); -lean_inc(x_477); -if (lean_is_exclusive(x_373)) { - lean_ctor_release(x_373, 0); - x_478 = x_373; +x_594 = lean_ctor_get(x_461, 0); +lean_inc(x_594); +if (lean_is_exclusive(x_461)) { + lean_ctor_release(x_461, 0); + x_595 = x_461; } else { - lean_dec_ref(x_373); - x_478 = lean_box(0); + lean_dec_ref(x_461); + x_595 = lean_box(0); } -if (lean_is_scalar(x_478)) { - x_479 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_595)) { + x_596 = lean_alloc_ctor(1, 1, 0); } else { - x_479 = x_478; + x_596 = x_595; } -lean_ctor_set(x_479, 0, x_477); -return x_479; +lean_ctor_set(x_596, 0, x_594); +return x_596; } } else { -lean_object* x_480; lean_object* x_481; lean_object* x_482; -lean_dec(x_361); -lean_dec(x_360); -lean_dec(x_359); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_597; lean_object* x_598; lean_object* x_599; +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_447); +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_2); -x_480 = lean_ctor_get(x_371, 0); -lean_inc(x_480); -if (lean_is_exclusive(x_371)) { - lean_ctor_release(x_371, 0); - x_481 = x_371; +x_597 = lean_ctor_get(x_459, 0); +lean_inc(x_597); +if (lean_is_exclusive(x_459)) { + lean_ctor_release(x_459, 0); + x_598 = x_459; } else { - lean_dec_ref(x_371); - x_481 = lean_box(0); + lean_dec_ref(x_459); + x_598 = lean_box(0); } -if (lean_is_scalar(x_481)) { - x_482 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_598)) { + x_599 = lean_alloc_ctor(1, 1, 0); } else { - x_482 = x_481; + x_599 = x_598; } -lean_ctor_set(x_482, 0, x_480); -return x_482; +lean_ctor_set(x_599, 0, x_597); +return x_599; } } else { -lean_object* x_483; lean_object* x_484; lean_object* x_485; -lean_dec(x_361); -lean_dec(x_360); -lean_dec(x_359); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_600; lean_object* x_601; lean_object* x_602; +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_447); +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_2); -x_483 = lean_ctor_get(x_369, 0); -lean_inc(x_483); -if (lean_is_exclusive(x_369)) { - lean_ctor_release(x_369, 0); - x_484 = x_369; +x_600 = lean_ctor_get(x_457, 0); +lean_inc(x_600); +if (lean_is_exclusive(x_457)) { + lean_ctor_release(x_457, 0); + x_601 = x_457; } else { - lean_dec_ref(x_369); - x_484 = lean_box(0); + lean_dec_ref(x_457); + x_601 = lean_box(0); } -if (lean_is_scalar(x_484)) { - x_485 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_601)) { + x_602 = lean_alloc_ctor(1, 1, 0); } else { - x_485 = x_484; + x_602 = x_601; } -lean_ctor_set(x_485, 0, x_483); -return x_485; +lean_ctor_set(x_602, 0, x_600); +return x_602; } } else { -lean_object* x_486; lean_object* x_487; lean_object* x_488; -lean_dec(x_361); -lean_dec(x_360); -lean_dec(x_359); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_6); +lean_object* x_603; lean_object* x_604; lean_object* x_605; +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_447); +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_2); -x_486 = lean_ctor_get(x_367, 0); -lean_inc(x_486); -if (lean_is_exclusive(x_367)) { - lean_ctor_release(x_367, 0); - x_487 = x_367; +x_603 = lean_ctor_get(x_455, 0); +lean_inc(x_603); +if (lean_is_exclusive(x_455)) { + lean_ctor_release(x_455, 0); + x_604 = x_455; } else { - lean_dec_ref(x_367); - x_487 = lean_box(0); + lean_dec_ref(x_455); + x_604 = lean_box(0); } -if (lean_is_scalar(x_487)) { - x_488 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_604)) { + x_605 = lean_alloc_ctor(1, 1, 0); } else { - x_488 = x_487; + x_605 = x_604; } -lean_ctor_set(x_488, 0, x_486); -return x_488; +lean_ctor_set(x_605, 0, x_603); +return x_605; } } } } } } -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; -x_14 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_1, x_2, x_5, x_6, x_9, x_10, x_11, x_12); -return x_14; +lean_object* x_15; +x_15 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_1, x_2, x_3, x_6, x_7, x_10, x_11, x_12, x_13); +return x_15; } } 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) { @@ -8578,14 +9143,16 @@ goto block_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; +uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get_uint8(x_7, sizeof(void*)*1); 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; +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_4, x_22); +x_24 = lean_alloc_ctor(2, 1, 1); +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_21); +x_10 = x_24; +x_11 = x_23; goto block_16; } } @@ -8684,7 +9251,7 @@ return x_6; 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, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_115; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t 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_46; uint8_t x_47; uint8_t 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_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_115; x_59 = lean_mk_empty_array_with_capacity(x_4); x_60 = lean_unsigned_to_nat(0u); x_110 = lean_nat_add(x_6, x_4); @@ -8719,11 +9286,11 @@ return x_21; } block_33: { -if (x_25 == 0) +if (x_24 == 0) { lean_dec_ref(x_3); x_14 = x_23; -x_15 = x_24; +x_15 = x_25; x_16 = x_26; x_17 = x_27; x_18 = lean_box(0); @@ -8740,7 +9307,7 @@ x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec_ref(x_30); x_14 = x_23; -x_15 = x_24; +x_15 = x_25; x_16 = x_31; x_17 = x_32; x_18 = lean_box(0); @@ -8749,10 +9316,10 @@ goto block_22; } block_45: { -if (x_34 == 0) +if (x_36 == 0) { x_23 = x_35; -x_24 = x_36; +x_24 = x_34; x_25 = x_37; x_26 = x_38; x_27 = x_39; @@ -8771,7 +9338,7 @@ x_44 = lean_ctor_get(x_42, 1); lean_inc(x_44); lean_dec_ref(x_42); x_23 = x_35; -x_24 = x_36; +x_24 = x_34; x_25 = x_37; x_26 = x_43; x_27 = x_44; @@ -8781,11 +9348,11 @@ goto block_33; } block_58: { -if (x_49 == 0) +if (x_48 == 0) { -x_34 = x_46; -x_35 = x_47; -x_36 = x_48; +x_34 = x_47; +x_35 = x_46; +x_36 = x_49; x_37 = x_50; x_38 = x_51; x_39 = x_52; @@ -8803,9 +9370,9 @@ lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); lean_dec_ref(x_55); -x_34 = x_46; -x_35 = x_47; -x_36 = x_48; +x_34 = x_47; +x_35 = x_46; +x_36 = x_49; x_37 = x_50; x_38 = x_56; x_39 = x_57; @@ -8845,7 +9412,7 @@ lean_ctor_set(x_75, 0, x_64); x_76 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_76, 0, x_75); lean_ctor_set(x_76, 1, x_65); -x_77 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_63, x_60, x_76, x_74, x_9, x_10, x_11, x_12); +x_77 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_63, x_60, x_5, x_76, x_74, x_9, x_10, x_11, x_12); lean_dec_ref(x_63); if (lean_obj_tag(x_77) == 0) { @@ -8894,17 +9461,17 @@ lean_dec(x_87); if (x_95 == 0) { uint8_t x_96; uint8_t x_97; uint8_t x_98; -x_96 = lean_unbox(x_85); -lean_dec(x_85); +x_96 = lean_unbox(x_84); +lean_dec(x_84); x_97 = lean_unbox(x_86); lean_dec(x_86); -x_98 = lean_unbox(x_84); -lean_dec(x_84); -x_46 = x_96; -x_47 = x_88; -x_48 = x_94; -x_49 = x_97; -x_50 = x_98; +x_98 = lean_unbox(x_85); +lean_dec(x_85); +x_46 = x_88; +x_47 = x_96; +x_48 = x_97; +x_49 = x_98; +x_50 = x_94; x_51 = x_92; x_52 = x_60; x_53 = lean_box(0); @@ -8921,17 +9488,17 @@ lean_inc(x_101); x_102 = lean_ctor_get(x_100, 1); lean_inc(x_102); lean_dec_ref(x_100); -x_103 = lean_unbox(x_85); -lean_dec(x_85); +x_103 = lean_unbox(x_84); +lean_dec(x_84); x_104 = lean_unbox(x_86); lean_dec(x_86); -x_105 = lean_unbox(x_84); -lean_dec(x_84); -x_46 = x_103; -x_47 = x_88; -x_48 = x_94; -x_49 = x_104; -x_50 = x_105; +x_105 = lean_unbox(x_85); +lean_dec(x_85); +x_46 = x_88; +x_47 = x_103; +x_48 = x_104; +x_49 = x_105; +x_50 = x_94; x_51 = x_101; x_52 = x_102; x_53 = lean_box(0); @@ -8988,7 +9555,7 @@ static lean_object* _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCt lean_object* 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___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__2; x_2 = lean_unsigned_to_nat(64u); -x_3 = lean_unsigned_to_nat(169u); +x_3 = lean_unsigned_to_nat(174u); x_4 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg___closed__0; x_5 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache_spec__3_spec__3___redArg___closed__0; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -9241,22 +9808,24 @@ lean_dec_ref(x_8); return x_17; } } -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, x_8); +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +x_12 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3___redArg(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec_ref(x_1); -return x_10; +return x_12; } } -LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___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, x_11, x_12); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_3); +x_16 = l_Std_Iterators_IterM_DefaultConsumers_forIn_x27___at___00__private_Lean_Compiler_IR_ToIRType_0__Lean_IR_getCtorLayout_fillCache_spec__3(x_1, x_2, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec_ref(x_1); -return x_14; +return x_16; } } 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) { @@ -9594,6 +10163,16 @@ l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__26); l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__27 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__27(); lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__27); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__28); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__29); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__30); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__31); +l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__32 = _init_l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__32(); +lean_mark_persistent(l___private_Lean_Compiler_IR_ToIRType_0__Lean_IR_nameToIRType_fillCache___closed__32); l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0 = _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0(); lean_mark_persistent(l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__0); l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1 = _init_l_Lean_Compiler_LCNF_CacheExtension_find_x3f___at___00Lean_IR_nameToIRType_spec__0___redArg___closed__1(); @@ -9650,6 +10229,10 @@ l_Lean_IR_CtorFieldInfo_format___closed__12 = _init_l_Lean_IR_CtorFieldInfo_form 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_format___closed__14 = _init_l_Lean_IR_CtorFieldInfo_format___closed__14(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__14); +l_Lean_IR_CtorFieldInfo_format___closed__15 = _init_l_Lean_IR_CtorFieldInfo_format___closed__15(); +lean_mark_persistent(l_Lean_IR_CtorFieldInfo_format___closed__15); 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(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c index d7908ef0c5fd..edcece022bb9 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c @@ -27,6 +27,7 @@ static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__28; static lean_object* l_Lean_Compiler_LCNF_casesThunkToMono___redArg___closed__0; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_argToMonoDeferredCheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___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_Compiler_LCNF_trivialStructToMono___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_toMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__24_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; @@ -52,9 +53,11 @@ static lean_object* l_Lean_Compiler_LCNF_casesThunkToMono___redArg___closed__5; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__1_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__7; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__3_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__50; extern lean_object* l_Lean_instEmptyCollectionFVarIdHashSet; static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__15; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__5_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__10; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4(lean_object*, lean_object*, 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_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_ctorAppToMono_spec__0___redArg(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -68,6 +71,7 @@ static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__17; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesThunkToMono___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__44; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__20; static lean_object* l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__0; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___00Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_Compiler_LCNF_checkFVarUseDeferred_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); @@ -95,6 +99,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesTaskToMono___redArg(lean_obje lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_Compiler_LCNF_checkFVarUse_spec__0(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__20; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__49; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_Const_get_x3f___at___00Lean_Compiler_LCNF_checkFVarUse_spec__0___redArg(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_throwNoncomputableError___redArg___closed__4; @@ -112,6 +117,7 @@ LEAN_EXPORT lean_object* l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__3 static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__2_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_argsToMonoRedArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__41; static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; 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_Std_DHashMap_Internal_Raw_u2080_insertIfNew___at___00Lean_Compiler_LCNF_Param_toMono_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); @@ -139,6 +145,7 @@ static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_ uint8_t lean_string_dec_eq(lean_object*, lean_object*); static 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_Compiler_LCNF_Code_toMono_spec__4_spec__4___closed__0; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesByteArrayToMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__48; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__19_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesUIntToMono___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_Compiler_LCNF_mkFieldParamsForComputedFields_spec__1___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,6 +171,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_argToMono(lean_object*, lean_objec size_t lean_ptr_addr(lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_casesIntToMono_spec__0___closed__1; static lean_object* l_Lean_Compiler_LCNF_casesIntToMono___redArg___closed__0; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__46; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__21_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__18; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToMonoM_State_ctorIdx(lean_object*); @@ -171,6 +179,7 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_casesNatToMono_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*); lean_object* l_Lean_Compiler_LCNF_Decl_saveMono___redArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__22_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; +static lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_FunDecl_toMono_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___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,11 +291,14 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_decToMono___boxed(lean_object*, le extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam_default; static lean_object* l_Lean_Compiler_LCNF_Decl_toMono___closed__5; uint8_t l_Lean_isExtern(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2; lean_object* l_Lean_Compiler_LCNF_LCtx_addFunDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_Compiler_LCNF_mkFieldParamsForComputedFields_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__35; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__39; LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_Compiler_LCNF_mkFieldParamsForComputedFields_spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_insert___at___00Lean_Compiler_LCNF_checkFVarUseDeferred_spec__0___redArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__47; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ctorAppToMono___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_Compiler_LCNF_FunDecl_toMono_spec__0(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__3; @@ -325,6 +337,7 @@ static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__22; static lean_object* l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__4; static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__0; static lean_object* l_Lean_Compiler_LCNF_casesThunkToMono___redArg___closed__6; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__43; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__36; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__18; @@ -343,6 +356,8 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_instBEqFVarId_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_argToMonoBase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__45; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__40; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__7; static lean_object* l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__0; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Compiler_LCNF_argsToMonoWithFnType_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*); @@ -352,6 +367,7 @@ lean_object* l_Lean_Compiler_LCNF_toMonoType(lean_object*, lean_object*, lean_ob static lean_object* l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__2; LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Lean_Compiler_LCNF_mkFieldParamsForComputedFields_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__23; +static lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_ToMono_0__Lean_Compiler_LCNF_initFn___closed__15_00___x40_Lean_Compiler_LCNF_ToMono_1770774466____hygCtx___hyg_2_; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___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_Std_DHashMap_Internal_Raw_u2080_contains___at___00Lean_Compiler_LCNF_argToMono_spec__0___boxed(lean_object*, lean_object*, lean_object*); @@ -364,6 +380,7 @@ lean_object* l_Lean_Compiler_LCNF_mkParam(lean_object*, lean_object*, uint8_t, l extern lean_object* l_Lean_Compiler_LCNF_instInhabitedCode_default__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__0; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__2; lean_object* l_Lean_Name_mkStr2(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_insertIfNew___at___00Lean_Compiler_LCNF_Param_toMono_spec__0_spec__1_spec__1_spec__1(lean_object*, lean_object*, lean_object*); @@ -422,12 +439,14 @@ static lean_object* l_Lean_Compiler_LCNF_casesNatToMono___redArg___closed__2; size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__0; static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__12; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_LetValue_toMono___closed__31; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDecl_toMono___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesIntToMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__5; static lean_object* l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0; static lean_object* l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__4; +static lean_object* l_Lean_Compiler_LCNF_Code_toMono___closed__42; 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_Compiler_LCNF_argsToMonoRedArg_spec__1_spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltCodeImp(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -7952,7 +7971,7 @@ static lean_object* _init_l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_ lean_object* 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_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4_spec__4___closed__1; x_2 = lean_unsigned_to_nat(70u); -x_3 = lean_unsigned_to_nat(416u); +x_3 = lean_unsigned_to_nat(435u); x_4 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4_spec__4___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -9109,7 +9128,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__1() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_Code_toMono___closed__0; x_2 = lean_unsigned_to_nat(66u); -x_3 = lean_unsigned_to_nat(406u); +x_3 = lean_unsigned_to_nat(425u); x_4 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4_spec__4___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -9151,7 +9170,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(27u); -x_3 = lean_unsigned_to_nat(361u); +x_3 = lean_unsigned_to_nat(372u); x_4 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4_spec__4___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -9329,7 +9348,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__25() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Array", 5, 5); +x_1 = lean_mk_string_unchecked("Int8", 4, 4); return x_1; } } @@ -9346,7 +9365,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__27() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("ByteArray", 9, 9); +x_1 = lean_mk_string_unchecked("Int16", 5, 5); return x_1; } } @@ -9363,7 +9382,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__29() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("FloatArray", 10, 10); +x_1 = lean_mk_string_unchecked("Int32", 5, 5); return x_1; } } @@ -9380,7 +9399,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__31() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("String", 6, 6); +x_1 = lean_mk_string_unchecked("Int64", 5, 5); return x_1; } } @@ -9397,7 +9416,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__33() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Thunk", 5, 5); +x_1 = lean_mk_string_unchecked("Array", 5, 5); return x_1; } } @@ -9414,7 +9433,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__35() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Task", 4, 4); +x_1 = lean_mk_string_unchecked("ByteArray", 9, 9); return x_1; } } @@ -9431,7 +9450,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__37() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_impl", 5, 5); +x_1 = lean_mk_string_unchecked("FloatArray", 10, 10); return x_1; } } @@ -9444,6 +9463,106 @@ x_2 = l_Lean_Name_mkStr1(x_1); return x_2; } } +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__39() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("String", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_Code_toMono___closed__39; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__41() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Thunk", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__42() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_Code_toMono___closed__41; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__43() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Task", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__44() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_Code_toMono___closed__43; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__45() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_impl", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__46() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_Code_toMono___closed__45; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__47() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toUInt64", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__48() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toUInt32", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__49() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toUInt16", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_Code_toMono___closed__50() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toUInt8", 7, 7); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_toMono(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -10146,8 +10265,8 @@ x_94 = lean_ptr_addr(x_92); x_95 = lean_usize_dec_eq(x_93, x_94); if (x_95 == 0) { -x_17 = lean_box(0); -x_18 = x_90; +x_17 = x_90; +x_18 = lean_box(0); x_19 = x_92; x_20 = x_95; goto block_24; @@ -10158,8 +10277,8 @@ size_t x_96; size_t x_97; uint8_t x_98; x_96 = lean_ptr_addr(x_81); x_97 = lean_ptr_addr(x_90); x_98 = lean_usize_dec_eq(x_96, x_97); -x_17 = lean_box(0); -x_18 = x_90; +x_17 = x_90; +x_18 = lean_box(0); x_19 = x_92; x_20 = x_98; goto block_24; @@ -10392,20 +10511,40 @@ x_281 = l_Lean_Compiler_LCNF_Code_toMono___closed__36; x_282 = lean_name_eq(x_252, x_281); if (x_282 == 0) { -lean_object* x_283; +lean_object* x_283; uint8_t x_284; +x_283 = l_Lean_Compiler_LCNF_Code_toMono___closed__38; +x_284 = lean_name_eq(x_252, x_283); +if (x_284 == 0) +{ +lean_object* x_285; uint8_t x_286; +x_285 = l_Lean_Compiler_LCNF_Code_toMono___closed__40; +x_286 = lean_name_eq(x_252, x_285); +if (x_286 == 0) +{ +lean_object* x_287; uint8_t x_288; +x_287 = l_Lean_Compiler_LCNF_Code_toMono___closed__42; +x_288 = lean_name_eq(x_252, x_287); +if (x_288 == 0) +{ +lean_object* x_289; uint8_t x_290; +x_289 = l_Lean_Compiler_LCNF_Code_toMono___closed__44; +x_290 = lean_name_eq(x_252, x_289); +if (x_290 == 0) +{ +lean_object* x_291; lean_inc(x_6); lean_inc_ref(x_5); lean_inc(x_252); -x_283 = l_Lean_Compiler_LCNF_hasTrivialStructure_x3f(x_252, x_5, x_6); -if (lean_obj_tag(x_283) == 0) +x_291 = l_Lean_Compiler_LCNF_hasTrivialStructure_x3f(x_252, x_5, x_6); +if (lean_obj_tag(x_291) == 0) { -lean_object* x_284; -x_284 = lean_ctor_get(x_283, 0); -lean_inc(x_284); -lean_dec_ref(x_283); -if (lean_obj_tag(x_284) == 0) +lean_object* x_292; +x_292 = lean_ctor_get(x_291, 0); +lean_inc(x_292); +lean_dec_ref(x_291); +if (lean_obj_tag(x_292) == 0) { -lean_object* x_285; lean_object* x_286; +lean_object* x_293; lean_object* x_294; lean_inc_ref(x_255); lean_inc(x_254); lean_inc_ref(x_253); @@ -10415,40 +10554,40 @@ if (lean_is_exclusive(x_251)) { lean_ctor_release(x_251, 1); lean_ctor_release(x_251, 2); lean_ctor_release(x_251, 3); - x_285 = x_251; + x_293 = x_251; } else { lean_dec_ref(x_251); - x_285 = lean_box(0); + x_293 = lean_box(0); } lean_inc(x_6); lean_inc_ref(x_5); lean_inc_ref(x_253); -x_286 = l_Lean_Compiler_LCNF_toMonoType(x_253, x_5, x_6); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_309; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - x_288 = x_286; -} else { - lean_dec_ref(x_286); - x_288 = lean_box(0); -} -x_289 = lean_st_ref_get(x_6); -x_290 = lean_ctor_get(x_289, 0); -lean_inc_ref(x_290); -lean_dec_ref(x_289); +x_294 = l_Lean_Compiler_LCNF_toMonoType(x_253, x_5, x_6); +if (lean_obj_tag(x_294) == 0) +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_317; +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + x_296 = x_294; +} else { + lean_dec_ref(x_294); + x_296 = lean_box(0); +} +x_297 = lean_st_ref_get(x_6); +x_298 = lean_ctor_get(x_297, 0); +lean_inc_ref(x_298); +lean_dec_ref(x_297); lean_inc(x_252); -lean_inc_ref(x_290); -x_309 = l_Lean_Environment_find_x3f(x_290, x_252, x_282); -if (lean_obj_tag(x_309) == 0) -{ -lean_dec_ref(x_290); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_285); +lean_inc_ref(x_298); +x_317 = l_Lean_Environment_find_x3f(x_298, x_252, x_290); +if (lean_obj_tag(x_317) == 0) +{ +lean_dec_ref(x_298); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_293); lean_dec_ref(x_255); lean_dec(x_254); lean_dec_ref(x_253); @@ -10464,189 +10603,189 @@ goto block_16; } else { -lean_object* x_310; -x_310 = lean_ctor_get(x_309, 0); -lean_inc(x_310); -lean_dec_ref(x_309); -if (lean_obj_tag(x_310) == 5) -{ -lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -x_311 = lean_ctor_get(x_310, 0); -lean_inc_ref(x_311); -if (lean_is_exclusive(x_310)) { - lean_ctor_release(x_310, 0); - x_312 = x_310; -} else { - lean_dec_ref(x_310); - x_312 = lean_box(0); -} -x_313 = lean_ctor_get(x_311, 0); -lean_inc_ref(x_313); -lean_dec_ref(x_311); -x_314 = lean_ctor_get(x_313, 0); -lean_inc(x_314); -lean_dec_ref(x_313); -x_315 = l_Lean_mkCasesOnName(x_314); -lean_inc_ref(x_290); -x_316 = l_Lean_Compiler_getImplementedBy_x3f(x_290, x_315); -if (lean_obj_tag(x_316) == 0) -{ -if (x_282 == 0) -{ -size_t x_317; size_t x_318; lean_object* x_319; -lean_dec_ref(x_290); -lean_dec(x_285); -x_317 = lean_array_size(x_255); -x_318 = 0; +lean_object* x_318; +x_318 = lean_ctor_get(x_317, 0); +lean_inc(x_318); +lean_dec_ref(x_317); +if (lean_obj_tag(x_318) == 5) +{ +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +x_319 = lean_ctor_get(x_318, 0); +lean_inc_ref(x_319); +if (lean_is_exclusive(x_318)) { + lean_ctor_release(x_318, 0); + x_320 = x_318; +} else { + lean_dec_ref(x_318); + x_320 = lean_box(0); +} +x_321 = lean_ctor_get(x_319, 0); +lean_inc_ref(x_321); +lean_dec_ref(x_319); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +lean_dec_ref(x_321); +x_323 = l_Lean_mkCasesOnName(x_322); +lean_inc_ref(x_298); +x_324 = l_Lean_Compiler_getImplementedBy_x3f(x_298, x_323); +if (lean_obj_tag(x_324) == 0) +{ +if (x_290 == 0) +{ +size_t x_325; size_t x_326; lean_object* x_327; +lean_dec_ref(x_298); +lean_dec(x_293); +x_325 = lean_array_size(x_255); +x_326 = 0; lean_inc_ref(x_255); -x_319 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__6(x_317, x_318, x_255, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_319) == 0) +x_327 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__6(x_325, x_326, x_255, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_327) == 0) { -lean_object* x_320; lean_object* x_321; uint8_t x_326; size_t x_330; size_t x_331; uint8_t x_332; -x_320 = lean_ctor_get(x_319, 0); -lean_inc(x_320); -if (lean_is_exclusive(x_319)) { - lean_ctor_release(x_319, 0); - x_321 = x_319; +lean_object* x_328; lean_object* x_329; uint8_t x_334; size_t x_338; size_t x_339; uint8_t x_340; +x_328 = lean_ctor_get(x_327, 0); +lean_inc(x_328); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + x_329 = x_327; } else { - lean_dec_ref(x_319); - x_321 = lean_box(0); + lean_dec_ref(x_327); + x_329 = lean_box(0); } -x_330 = lean_ptr_addr(x_255); +x_338 = lean_ptr_addr(x_255); lean_dec_ref(x_255); -x_331 = lean_ptr_addr(x_320); -x_332 = lean_usize_dec_eq(x_330, x_331); -if (x_332 == 0) +x_339 = lean_ptr_addr(x_328); +x_340 = lean_usize_dec_eq(x_338, x_339); +if (x_340 == 0) { lean_dec_ref(x_253); -x_326 = x_332; -goto block_329; +x_334 = x_340; +goto block_337; } else { -size_t x_333; size_t x_334; uint8_t x_335; -x_333 = lean_ptr_addr(x_253); +size_t x_341; size_t x_342; uint8_t x_343; +x_341 = lean_ptr_addr(x_253); lean_dec_ref(x_253); -x_334 = lean_ptr_addr(x_287); -x_335 = lean_usize_dec_eq(x_333, x_334); -x_326 = x_335; -goto block_329; +x_342 = lean_ptr_addr(x_295); +x_343 = lean_usize_dec_eq(x_341, x_342); +x_334 = x_343; +goto block_337; } -block_325: +block_333: { -lean_object* x_322; lean_object* x_323; lean_object* x_324; -x_322 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_322, 0, x_252); -lean_ctor_set(x_322, 1, x_287); -lean_ctor_set(x_322, 2, x_254); -lean_ctor_set(x_322, 3, x_320); -if (lean_is_scalar(x_312)) { - x_323 = lean_alloc_ctor(4, 1, 0); +lean_object* x_330; lean_object* x_331; lean_object* x_332; +x_330 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_330, 0, x_252); +lean_ctor_set(x_330, 1, x_295); +lean_ctor_set(x_330, 2, x_254); +lean_ctor_set(x_330, 3, x_328); +if (lean_is_scalar(x_320)) { + x_331 = lean_alloc_ctor(4, 1, 0); } else { - x_323 = x_312; - lean_ctor_set_tag(x_323, 4); + x_331 = x_320; + lean_ctor_set_tag(x_331, 4); } -lean_ctor_set(x_323, 0, x_322); -if (lean_is_scalar(x_321)) { - x_324 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_331, 0, x_330); +if (lean_is_scalar(x_329)) { + x_332 = lean_alloc_ctor(0, 1, 0); } else { - x_324 = x_321; + x_332 = x_329; } -lean_ctor_set(x_324, 0, x_323); -return x_324; +lean_ctor_set(x_332, 0, x_331); +return x_332; } -block_329: +block_337: { -if (x_326 == 0) +if (x_334 == 0) { -lean_dec(x_288); +lean_dec(x_296); lean_dec_ref(x_1); -goto block_325; +goto block_333; } else { -uint8_t x_327; -x_327 = l_Lean_instBEqFVarId_beq(x_254, x_254); -if (x_327 == 0) +uint8_t x_335; +x_335 = l_Lean_instBEqFVarId_beq(x_254, x_254); +if (x_335 == 0) { -lean_dec(x_288); +lean_dec(x_296); lean_dec_ref(x_1); -goto block_325; +goto block_333; } else { -lean_object* x_328; -lean_dec(x_321); +lean_object* x_336; +lean_dec(x_329); +lean_dec(x_328); lean_dec(x_320); -lean_dec(x_312); -lean_dec(x_287); +lean_dec(x_295); lean_dec(x_254); lean_dec(x_252); -if (lean_is_scalar(x_288)) { - x_328 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_296)) { + x_336 = lean_alloc_ctor(0, 1, 0); } else { - x_328 = x_288; + x_336 = x_296; } -lean_ctor_set(x_328, 0, x_1); -return x_328; +lean_ctor_set(x_336, 0, x_1); +return x_336; } } } } else { -uint8_t x_336; -lean_dec(x_312); -lean_dec(x_288); -lean_dec(x_287); +uint8_t x_344; +lean_dec(x_320); +lean_dec(x_296); +lean_dec(x_295); lean_dec_ref(x_255); lean_dec(x_254); lean_dec_ref(x_253); lean_dec(x_252); lean_dec_ref(x_1); -x_336 = !lean_is_exclusive(x_319); -if (x_336 == 0) +x_344 = !lean_is_exclusive(x_327); +if (x_344 == 0) { -return x_319; +return x_327; } else { -lean_object* x_337; lean_object* x_338; -x_337 = lean_ctor_get(x_319, 0); -lean_inc(x_337); -lean_dec(x_319); -x_338 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_338, 0, x_337); -return x_338; +lean_object* x_345; lean_object* x_346; +x_345 = lean_ctor_get(x_327, 0); +lean_inc(x_345); +lean_dec(x_327); +x_346 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_346, 0, x_345); +return x_346; } } } else { -lean_dec(x_312); -lean_dec(x_288); +lean_dec(x_320); +lean_dec(x_296); lean_dec_ref(x_253); lean_dec_ref(x_1); -goto block_308; +goto block_316; } } else { -lean_dec_ref(x_316); -lean_dec(x_312); -lean_dec(x_288); +lean_dec_ref(x_324); +lean_dec(x_320); +lean_dec(x_296); lean_dec_ref(x_253); lean_dec_ref(x_1); -goto block_308; +goto block_316; } } else { -lean_dec(x_310); -lean_dec_ref(x_290); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_285); +lean_dec(x_318); +lean_dec_ref(x_298); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_293); lean_dec_ref(x_255); lean_dec(x_254); lean_dec_ref(x_253); @@ -10661,88 +10800,88 @@ x_13 = lean_box(0); goto block_16; } } -block_308: +block_316: { -lean_object* x_291; size_t x_292; size_t x_293; lean_object* x_294; -x_291 = l_Lean_Compiler_LCNF_Code_toMono___closed__38; -x_292 = lean_array_size(x_255); -x_293 = 0; -x_294 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4(x_291, x_290, x_282, x_292, x_293, x_255, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_294) == 0) +lean_object* x_299; size_t x_300; size_t x_301; lean_object* x_302; +x_299 = l_Lean_Compiler_LCNF_Code_toMono___closed__46; +x_300 = lean_array_size(x_255); +x_301 = 0; +x_302 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_Code_toMono_spec__4(x_299, x_298, x_290, x_300, x_301, x_255, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_302) == 0) { -uint8_t x_295; -x_295 = !lean_is_exclusive(x_294); -if (x_295 == 0) -{ -lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; -x_296 = lean_ctor_get(x_294, 0); -x_297 = l_Lean_Name_append(x_252, x_291); -if (lean_is_scalar(x_285)) { - x_298 = lean_alloc_ctor(0, 4, 0); -} else { - x_298 = x_285; -} -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_287); -lean_ctor_set(x_298, 2, x_254); -lean_ctor_set(x_298, 3, x_296); -x_299 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_294, 0, x_299); -return x_294; +uint8_t x_303; +x_303 = !lean_is_exclusive(x_302); +if (x_303 == 0) +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +x_304 = lean_ctor_get(x_302, 0); +x_305 = l_Lean_Name_append(x_252, x_299); +if (lean_is_scalar(x_293)) { + x_306 = lean_alloc_ctor(0, 4, 0); +} else { + x_306 = x_293; +} +lean_ctor_set(x_306, 0, x_305); +lean_ctor_set(x_306, 1, x_295); +lean_ctor_set(x_306, 2, x_254); +lean_ctor_set(x_306, 3, x_304); +x_307 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_307, 0, x_306); +lean_ctor_set(x_302, 0, x_307); +return x_302; } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; -x_300 = lean_ctor_get(x_294, 0); -lean_inc(x_300); -lean_dec(x_294); -x_301 = l_Lean_Name_append(x_252, x_291); -if (lean_is_scalar(x_285)) { - x_302 = lean_alloc_ctor(0, 4, 0); +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_308 = lean_ctor_get(x_302, 0); +lean_inc(x_308); +lean_dec(x_302); +x_309 = l_Lean_Name_append(x_252, x_299); +if (lean_is_scalar(x_293)) { + x_310 = lean_alloc_ctor(0, 4, 0); } else { - x_302 = x_285; + x_310 = x_293; } -lean_ctor_set(x_302, 0, x_301); -lean_ctor_set(x_302, 1, x_287); -lean_ctor_set(x_302, 2, x_254); -lean_ctor_set(x_302, 3, x_300); -x_303 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_303, 0, x_302); -x_304 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_304, 0, x_303); -return x_304; +lean_ctor_set(x_310, 0, x_309); +lean_ctor_set(x_310, 1, x_295); +lean_ctor_set(x_310, 2, x_254); +lean_ctor_set(x_310, 3, x_308); +x_311 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_311, 0, x_310); +x_312 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_312, 0, x_311); +return x_312; } } else { -uint8_t x_305; -lean_dec(x_287); -lean_dec(x_285); +uint8_t x_313; +lean_dec(x_295); +lean_dec(x_293); lean_dec(x_254); lean_dec(x_252); -x_305 = !lean_is_exclusive(x_294); -if (x_305 == 0) +x_313 = !lean_is_exclusive(x_302); +if (x_313 == 0) { -return x_294; +return x_302; } else { -lean_object* x_306; lean_object* x_307; -x_306 = lean_ctor_get(x_294, 0); -lean_inc(x_306); -lean_dec(x_294); -x_307 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_307, 0, x_306); -return x_307; +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_302, 0); +lean_inc(x_314); +lean_dec(x_302); +x_315 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_314); +return x_315; } } } } else { -uint8_t x_339; -lean_dec(x_285); +uint8_t x_347; +lean_dec(x_293); lean_dec_ref(x_255); lean_dec(x_254); lean_dec_ref(x_253); @@ -10753,38 +10892,38 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); lean_dec_ref(x_1); -x_339 = !lean_is_exclusive(x_286); -if (x_339 == 0) +x_347 = !lean_is_exclusive(x_294); +if (x_347 == 0) { -return x_286; +return x_294; } else { -lean_object* x_340; lean_object* x_341; -x_340 = lean_ctor_get(x_286, 0); -lean_inc(x_340); -lean_dec(x_286); -x_341 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_341, 0, x_340); -return x_341; +lean_object* x_348; lean_object* x_349; +x_348 = lean_ctor_get(x_294, 0); +lean_inc(x_348); +lean_dec(x_294); +x_349 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_349, 0, x_348); +return x_349; } } } else { -lean_object* x_342; lean_object* x_343; +lean_object* x_350; lean_object* x_351; lean_dec_ref(x_1); -x_342 = lean_ctor_get(x_284, 0); -lean_inc(x_342); -lean_dec_ref(x_284); -x_343 = l_Lean_Compiler_LCNF_trivialStructToMono(x_342, x_251, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_342); -return x_343; +x_350 = lean_ctor_get(x_292, 0); +lean_inc(x_350); +lean_dec_ref(x_292); +x_351 = l_Lean_Compiler_LCNF_trivialStructToMono(x_350, x_251, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_350); +return x_351; } } else { -uint8_t x_344; +uint8_t x_352; lean_dec_ref(x_251); lean_dec(x_6); lean_dec_ref(x_5); @@ -10792,131 +10931,167 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); lean_dec_ref(x_1); -x_344 = !lean_is_exclusive(x_283); -if (x_344 == 0) +x_352 = !lean_is_exclusive(x_291); +if (x_352 == 0) { -return x_283; +return x_291; } else { -lean_object* x_345; lean_object* x_346; -x_345 = lean_ctor_get(x_283, 0); -lean_inc(x_345); -lean_dec(x_283); -x_346 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_346, 0, x_345); -return x_346; +lean_object* x_353; lean_object* x_354; +x_353 = lean_ctor_get(x_291, 0); +lean_inc(x_353); +lean_dec(x_291); +x_354 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_354, 0, x_353); +return x_354; } } } else { -lean_object* x_347; +lean_object* x_355; lean_dec_ref(x_1); -x_347 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_347; +x_355 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_355; } } else { -lean_object* x_348; +lean_object* x_356; lean_dec_ref(x_1); -x_348 = l_Lean_Compiler_LCNF_casesThunkToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +x_356 = l_Lean_Compiler_LCNF_casesThunkToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); lean_dec_ref(x_251); -return x_348; +return x_356; } } else { -lean_object* x_349; +lean_object* x_357; lean_dec_ref(x_1); -x_349 = l_Lean_Compiler_LCNF_casesStringToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_349; +x_357 = l_Lean_Compiler_LCNF_casesStringToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_357; } } else { -lean_object* x_350; +lean_object* x_358; lean_dec_ref(x_1); -x_350 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_350; +x_358 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_358; } } else { -lean_object* x_351; +lean_object* x_359; lean_dec_ref(x_1); -x_351 = l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_351; +x_359 = l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_359; } } else { -lean_object* x_352; +lean_object* x_360; lean_dec_ref(x_1); -x_352 = l_Lean_Compiler_LCNF_casesArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_352; +x_360 = l_Lean_Compiler_LCNF_casesArrayToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_360; } } else { -lean_object* x_353; +lean_object* x_361; lean_object* x_362; lean_dec_ref(x_1); -x_353 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_269, x_2, x_3, x_4, x_5, x_6); -return x_353; +x_361 = l_Lean_Compiler_LCNF_Code_toMono___closed__47; +x_362 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg(x_251, x_277, x_361, x_2, x_3, x_4, x_5, x_6); +return x_362; } } else { -lean_object* x_354; +lean_object* x_363; lean_object* x_364; lean_dec_ref(x_1); -x_354 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_267, x_2, x_3, x_4, x_5, x_6); -return x_354; +x_363 = l_Lean_Compiler_LCNF_Code_toMono___closed__48; +x_364 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg(x_251, x_275, x_363, x_2, x_3, x_4, x_5, x_6); +return x_364; } } else { -lean_object* x_355; +lean_object* x_365; lean_object* x_366; lean_dec_ref(x_1); -x_355 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_265, x_2, x_3, x_4, x_5, x_6); -return x_355; +x_365 = l_Lean_Compiler_LCNF_Code_toMono___closed__49; +x_366 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg(x_251, x_273, x_365, x_2, x_3, x_4, x_5, x_6); +return x_366; } } else { -lean_object* x_356; +lean_object* x_367; lean_object* x_368; lean_dec_ref(x_1); -x_356 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_263, x_2, x_3, x_4, x_5, x_6); -return x_356; +x_367 = l_Lean_Compiler_LCNF_Code_toMono___closed__50; +x_368 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg(x_251, x_271, x_367, x_2, x_3, x_4, x_5, x_6); +return x_368; } } else { -lean_object* x_357; +lean_object* x_369; lean_dec_ref(x_1); -x_357 = l_Lean_Compiler_LCNF_casesIntToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_357; +x_369 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_269, x_2, x_3, x_4, x_5, x_6); +return x_369; } } else { -lean_object* x_358; +lean_object* x_370; lean_dec_ref(x_1); -x_358 = l_Lean_Compiler_LCNF_casesNatToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_358; +x_370 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_267, x_2, x_3, x_4, x_5, x_6); +return x_370; } } else { -lean_object* x_359; +lean_object* x_371; lean_dec_ref(x_1); -x_359 = l_Lean_Compiler_LCNF_decToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); -return x_359; +x_371 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_265, x_2, x_3, x_4, x_5, x_6); +return x_371; +} +} +else +{ +lean_object* x_372; +lean_dec_ref(x_1); +x_372 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg(x_251, x_263, x_2, x_3, x_4, x_5, x_6); +return x_372; +} +} +else +{ +lean_object* x_373; +lean_dec_ref(x_1); +x_373 = l_Lean_Compiler_LCNF_casesIntToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_373; +} +} +else +{ +lean_object* x_374; +lean_dec_ref(x_1); +x_374 = l_Lean_Compiler_LCNF_casesNatToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_374; +} +} +else +{ +lean_object* x_375; +lean_dec_ref(x_1); +x_375 = l_Lean_Compiler_LCNF_decToMono___redArg(x_251, x_2, x_3, x_4, x_5, x_6); +return x_375; } } else { -uint8_t x_360; +uint8_t x_376; lean_dec_ref(x_251); lean_dec(x_6); lean_dec_ref(x_5); @@ -10924,192 +11099,192 @@ lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); lean_dec_ref(x_1); -x_360 = !lean_is_exclusive(x_256); -if (x_360 == 0) +x_376 = !lean_is_exclusive(x_256); +if (x_376 == 0) { return x_256; } else { -lean_object* x_361; lean_object* x_362; -x_361 = lean_ctor_get(x_256, 0); -lean_inc(x_361); +lean_object* x_377; lean_object* x_378; +x_377 = lean_ctor_get(x_256, 0); +lean_inc(x_377); lean_dec(x_256); -x_362 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_362, 0, x_361); -return x_362; +x_378 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_378, 0, x_377); +return x_378; } } } case 5: { -lean_object* x_363; lean_object* x_364; -x_363 = lean_ctor_get(x_1, 0); -x_364 = l_Lean_Compiler_LCNF_checkFVarUse(x_363, x_2, x_3, x_4, x_5, x_6); +lean_object* x_379; lean_object* x_380; +x_379 = lean_ctor_get(x_1, 0); +x_380 = l_Lean_Compiler_LCNF_checkFVarUse(x_379, 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); -if (lean_obj_tag(x_364) == 0) +if (lean_obj_tag(x_380) == 0) { -uint8_t x_365; -x_365 = !lean_is_exclusive(x_364); -if (x_365 == 0) +uint8_t x_381; +x_381 = !lean_is_exclusive(x_380); +if (x_381 == 0) { -lean_object* x_366; -x_366 = lean_ctor_get(x_364, 0); -lean_dec(x_366); -lean_ctor_set(x_364, 0, x_1); -return x_364; +lean_object* x_382; +x_382 = lean_ctor_get(x_380, 0); +lean_dec(x_382); +lean_ctor_set(x_380, 0, x_1); +return x_380; } else { -lean_object* x_367; -lean_dec(x_364); -x_367 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_367, 0, x_1); -return x_367; +lean_object* x_383; +lean_dec(x_380); +x_383 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_383, 0, x_1); +return x_383; } } else { -uint8_t x_368; +uint8_t x_384; lean_dec_ref(x_1); -x_368 = !lean_is_exclusive(x_364); -if (x_368 == 0) +x_384 = !lean_is_exclusive(x_380); +if (x_384 == 0) { -return x_364; +return x_380; } else { -lean_object* x_369; lean_object* x_370; -x_369 = lean_ctor_get(x_364, 0); -lean_inc(x_369); -lean_dec(x_364); -x_370 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_370, 0, x_369); -return x_370; +lean_object* x_385; lean_object* x_386; +x_385 = lean_ctor_get(x_380, 0); +lean_inc(x_385); +lean_dec(x_380); +x_386 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_386, 0, x_385); +return x_386; } } } case 6: { -uint8_t x_371; +uint8_t x_387; lean_dec(x_4); lean_dec_ref(x_3); lean_dec(x_2); -x_371 = !lean_is_exclusive(x_1); -if (x_371 == 0) -{ -lean_object* x_372; lean_object* x_373; -x_372 = lean_ctor_get(x_1, 0); -x_373 = l_Lean_Compiler_LCNF_toMonoType(x_372, x_5, x_6); -if (lean_obj_tag(x_373) == 0) -{ -uint8_t x_374; -x_374 = !lean_is_exclusive(x_373); -if (x_374 == 0) -{ -lean_object* x_375; -x_375 = lean_ctor_get(x_373, 0); -lean_ctor_set(x_1, 0, x_375); -lean_ctor_set(x_373, 0, x_1); -return x_373; +x_387 = !lean_is_exclusive(x_1); +if (x_387 == 0) +{ +lean_object* x_388; lean_object* x_389; +x_388 = lean_ctor_get(x_1, 0); +x_389 = l_Lean_Compiler_LCNF_toMonoType(x_388, x_5, x_6); +if (lean_obj_tag(x_389) == 0) +{ +uint8_t x_390; +x_390 = !lean_is_exclusive(x_389); +if (x_390 == 0) +{ +lean_object* x_391; +x_391 = lean_ctor_get(x_389, 0); +lean_ctor_set(x_1, 0, x_391); +lean_ctor_set(x_389, 0, x_1); +return x_389; } else { -lean_object* x_376; lean_object* x_377; -x_376 = lean_ctor_get(x_373, 0); -lean_inc(x_376); -lean_dec(x_373); -lean_ctor_set(x_1, 0, x_376); -x_377 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_377, 0, x_1); -return x_377; +lean_object* x_392; lean_object* x_393; +x_392 = lean_ctor_get(x_389, 0); +lean_inc(x_392); +lean_dec(x_389); +lean_ctor_set(x_1, 0, x_392); +x_393 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_393, 0, x_1); +return x_393; } } else { -uint8_t x_378; +uint8_t x_394; lean_free_object(x_1); -x_378 = !lean_is_exclusive(x_373); -if (x_378 == 0) +x_394 = !lean_is_exclusive(x_389); +if (x_394 == 0) { -return x_373; +return x_389; } else { -lean_object* x_379; lean_object* x_380; -x_379 = lean_ctor_get(x_373, 0); -lean_inc(x_379); -lean_dec(x_373); -x_380 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_380, 0, x_379); -return x_380; +lean_object* x_395; lean_object* x_396; +x_395 = lean_ctor_get(x_389, 0); +lean_inc(x_395); +lean_dec(x_389); +x_396 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_396, 0, x_395); +return x_396; } } } else { -lean_object* x_381; lean_object* x_382; -x_381 = lean_ctor_get(x_1, 0); -lean_inc(x_381); +lean_object* x_397; lean_object* x_398; +x_397 = lean_ctor_get(x_1, 0); +lean_inc(x_397); lean_dec(x_1); -x_382 = l_Lean_Compiler_LCNF_toMonoType(x_381, x_5, x_6); -if (lean_obj_tag(x_382) == 0) +x_398 = l_Lean_Compiler_LCNF_toMonoType(x_397, x_5, x_6); +if (lean_obj_tag(x_398) == 0) { -lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; -x_383 = lean_ctor_get(x_382, 0); -lean_inc(x_383); -if (lean_is_exclusive(x_382)) { - lean_ctor_release(x_382, 0); - x_384 = x_382; +lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; +x_399 = lean_ctor_get(x_398, 0); +lean_inc(x_399); +if (lean_is_exclusive(x_398)) { + lean_ctor_release(x_398, 0); + x_400 = x_398; } else { - lean_dec_ref(x_382); - x_384 = lean_box(0); + lean_dec_ref(x_398); + x_400 = lean_box(0); } -x_385 = lean_alloc_ctor(6, 1, 0); -lean_ctor_set(x_385, 0, x_383); -if (lean_is_scalar(x_384)) { - x_386 = lean_alloc_ctor(0, 1, 0); +x_401 = lean_alloc_ctor(6, 1, 0); +lean_ctor_set(x_401, 0, x_399); +if (lean_is_scalar(x_400)) { + x_402 = lean_alloc_ctor(0, 1, 0); } else { - x_386 = x_384; + x_402 = x_400; } -lean_ctor_set(x_386, 0, x_385); -return x_386; +lean_ctor_set(x_402, 0, x_401); +return x_402; } else { -lean_object* x_387; lean_object* x_388; lean_object* x_389; -x_387 = lean_ctor_get(x_382, 0); -lean_inc(x_387); -if (lean_is_exclusive(x_382)) { - lean_ctor_release(x_382, 0); - x_388 = x_382; +lean_object* x_403; lean_object* x_404; lean_object* x_405; +x_403 = lean_ctor_get(x_398, 0); +lean_inc(x_403); +if (lean_is_exclusive(x_398)) { + lean_ctor_release(x_398, 0); + x_404 = x_398; } else { - lean_dec_ref(x_382); - x_388 = lean_box(0); + lean_dec_ref(x_398); + x_404 = lean_box(0); } -if (lean_is_scalar(x_388)) { - x_389 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_404)) { + x_405 = lean_alloc_ctor(1, 1, 0); } else { - x_389 = x_388; + x_405 = x_404; } -lean_ctor_set(x_389, 0, x_387); -return x_389; +lean_ctor_set(x_405, 0, x_403); +return x_405; } } } default: { -lean_object* x_390; lean_object* x_391; -x_390 = lean_ctor_get(x_1, 0); -x_391 = lean_ctor_get(x_1, 1); -lean_inc_ref(x_391); -lean_inc_ref(x_390); -x_41 = x_390; -x_42 = x_391; +lean_object* x_406; lean_object* x_407; +x_406 = lean_ctor_get(x_1, 0); +x_407 = lean_ctor_get(x_1, 1); +lean_inc_ref(x_407); +lean_inc_ref(x_406); +x_41 = x_406; +x_42 = x_407; x_43 = x_2; x_44 = x_3; x_45 = x_4; @@ -11133,7 +11308,7 @@ if (x_20 == 0) lean_object* x_21; lean_object* x_22; lean_dec_ref(x_1); x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 0, x_17); lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_22, 0, x_21); @@ -11143,7 +11318,7 @@ else { lean_object* x_23; lean_dec_ref(x_19); -lean_dec_ref(x_18); +lean_dec_ref(x_17); x_23 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_23, 0, x_1); return x_23; @@ -11156,8 +11331,8 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; lean_dec_ref(x_1); x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_26); -lean_ctor_set(x_29, 1, x_25); +lean_ctor_set(x_29, 0, x_25); +lean_ctor_set(x_29, 1, x_26); x_30 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_30, 0, x_29); return x_30; @@ -11179,8 +11354,8 @@ if (x_36 == 0) lean_object* x_37; lean_object* x_38; lean_dec_ref(x_1); x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 0, x_33); +lean_ctor_set(x_37, 1, x_34); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); return x_38; @@ -11227,8 +11402,8 @@ x_56 = lean_ptr_addr(x_52); x_57 = lean_usize_dec_eq(x_55, x_56); if (x_57 == 0) { -x_25 = x_52; -x_26 = x_50; +x_25 = x_50; +x_26 = x_52; x_27 = lean_box(0); x_28 = x_57; goto block_32; @@ -11239,8 +11414,8 @@ size_t x_58; size_t x_59; uint8_t x_60; x_58 = lean_ptr_addr(x_53); x_59 = lean_ptr_addr(x_50); x_60 = lean_usize_dec_eq(x_58, x_59); -x_25 = x_52; -x_26 = x_50; +x_25 = x_50; +x_26 = x_52; x_27 = lean_box(0); x_28 = x_60; goto block_32; @@ -11259,8 +11434,8 @@ x_65 = lean_ptr_addr(x_61); x_66 = lean_usize_dec_eq(x_64, x_65); if (x_66 == 0) { -x_33 = x_61; -x_34 = x_50; +x_33 = x_50; +x_34 = x_61; x_35 = lean_box(0); x_36 = x_66; goto block_40; @@ -11271,8 +11446,8 @@ size_t x_67; size_t x_68; uint8_t x_69; x_67 = lean_ptr_addr(x_62); x_68 = lean_ptr_addr(x_50); x_69 = lean_usize_dec_eq(x_67, x_68); -x_33 = x_61; -x_34 = x_50; +x_33 = x_50; +x_34 = x_61; x_35 = lean_box(0); x_36 = x_69; goto block_40; @@ -11365,7 +11540,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___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_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(344u); +x_3 = lean_unsigned_to_nat(355u); x_4 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11386,7 +11561,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__4() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__3; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(346u); +x_3 = lean_unsigned_to_nat(357u); x_4 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11407,7 +11582,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__6() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__5; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(347u); +x_3 = lean_unsigned_to_nat(358u); x_4 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11420,7 +11595,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_trivialStructToMono___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_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(41u); -x_3 = lean_unsigned_to_nat(345u); +x_3 = lean_unsigned_to_nat(356u); x_4 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11955,7 +12130,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(333u); +x_3 = lean_unsigned_to_nat(344u); x_4 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -11983,7 +12158,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__3; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__35; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__43; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -12004,7 +12179,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(334u); +x_3 = lean_unsigned_to_nat(345u); x_4 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -12548,7 +12723,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesThunkToMono___redArg___close lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(313u); +x_3 = lean_unsigned_to_nat(324u); x_4 = l_Lean_Compiler_LCNF_casesThunkToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -12560,7 +12735,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesThunkToMono___redArg___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__3; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__33; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__41; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -12607,7 +12782,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesThunkToMono___redArg___close lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(314u); +x_3 = lean_unsigned_to_nat(325u); x_4 = l_Lean_Compiler_LCNF_casesThunkToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -13277,7 +13452,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesStringToMono___redArg___clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(302u); +x_3 = lean_unsigned_to_nat(313u); x_4 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -13297,7 +13472,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesStringToMono___redArg___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__2; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__31; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__39; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -13317,7 +13492,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesStringToMono___redArg___clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(303u); +x_3 = lean_unsigned_to_nat(314u); x_4 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -13861,7 +14036,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesFloatArrayToMono___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; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(291u); +x_3 = lean_unsigned_to_nat(302u); x_4 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -13881,7 +14056,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__2; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__29; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__37; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -13892,7 +14067,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesFloatArrayToMono___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; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(292u); +x_3 = lean_unsigned_to_nat(303u); x_4 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -14436,7 +14611,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(280u); +x_3 = lean_unsigned_to_nat(291u); x_4 = l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -14448,7 +14623,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesFloatArrayToMono___redArg___closed__2; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__27; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__35; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -14459,7 +14634,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(281u); +x_3 = lean_unsigned_to_nat(292u); x_4 = l_Lean_Compiler_LCNF_casesByteArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -15003,7 +15178,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesArrayToMono___redArg___close lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(269u); +x_3 = lean_unsigned_to_nat(280u); x_4 = l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -15015,7 +15190,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesArrayToMono___redArg___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__2; -x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__25; +x_2 = l_Lean_Compiler_LCNF_Code_toMono___closed__33; x_3 = l_Lean_Name_mkStr2(x_2, x_1); return x_3; } @@ -15026,7 +15201,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_casesArrayToMono___redArg___close lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__21; x_2 = lean_unsigned_to_nat(34u); -x_3 = lean_unsigned_to_nat(270u); +x_3 = lean_unsigned_to_nat(281u); x_4 = l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); @@ -15556,34 +15731,605 @@ x_9 = l_Lean_Compiler_LCNF_casesArrayToMono___redArg(x_1, x_3, x_4, x_5, x_6, x_ return x_9; } } -static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0() { +static lean_object* _init_l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.casesUIntToMono", 34, 34); +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.casesSIntToMono", 34, 34); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_casesSIntToMono___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_Lean_Compiler_LCNF_trivialStructToMono___closed__1; x_2 = lean_unsigned_to_nat(2u); -x_3 = lean_unsigned_to_nat(258u); -x_4 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0; +x_3 = lean_unsigned_to_nat(269u); +x_4 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0; x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); return x_6; } } -static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("toBitVec", 8, 8); -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_Lean_Compiler_LCNF_LetValue_toMono___closed__21; +x_2 = lean_unsigned_to_nat(34u); +x_3 = lean_unsigned_to_nat(270u); +x_4 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0; +x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___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_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = lean_ctor_get(x_1, 1); +lean_dec(x_13); +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_get_size(x_12); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_1); +lean_dec_ref(x_12); +lean_dec(x_11); +lean_dec_ref(x_3); +lean_dec(x_2); +x_18 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1; +x_19 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__1(x_18, x_4, x_5, x_6, x_7, x_8); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__3___closed__0; +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_array_get(x_20, x_12, x_21); +lean_dec_ref(x_12); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_22, 1); +x_25 = lean_ctor_get(x_22, 2); +x_26 = lean_ctor_get(x_22, 0); +lean_dec(x_26); +x_27 = l_Lean_Compiler_LCNF_eraseParams___redArg(x_24, x_6); +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_dec_ref(x_27); +x_28 = lean_st_ref_take(x_6); +x_29 = l_Lean_Compiler_LCNF_argsToMonoRedArg___closed__0; +x_30 = lean_array_get(x_29, x_24, x_21); +lean_dec_ref(x_24); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec_ref(x_30); +x_33 = !lean_is_exclusive(x_28); +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 = lean_ctor_get(x_28, 0); +x_35 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__2; +x_36 = l_Lean_Name_str___override(x_2, x_3); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_11); +x_39 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__4; +x_40 = lean_array_push(x_39, x_38); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 2, x_40); +lean_ctor_set(x_22, 1, x_37); +lean_ctor_set(x_22, 0, x_36); +lean_ctor_set(x_1, 3, x_22); +lean_ctor_set(x_1, 2, x_35); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_31); +lean_inc_ref(x_1); +x_41 = l_Lean_Compiler_LCNF_LCtx_addLetDecl(x_34, x_1); +lean_ctor_set(x_28, 0, x_41); +x_42 = lean_st_ref_set(x_6, x_28); +x_43 = l_Lean_Compiler_LCNF_Code_toMono(x_25, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_1); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_43, 0, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_1); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_48); +return x_49; +} +} +else +{ +lean_dec_ref(x_1); +return x_43; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_50 = lean_ctor_get(x_28, 0); +x_51 = lean_ctor_get(x_28, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_28); +x_52 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__2; +x_53 = l_Lean_Name_str___override(x_2, x_3); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_11); +x_56 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__4; +x_57 = lean_array_push(x_56, x_55); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 2, x_57); +lean_ctor_set(x_22, 1, x_54); +lean_ctor_set(x_22, 0, x_53); +lean_ctor_set(x_1, 3, x_22); +lean_ctor_set(x_1, 2, x_52); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_31); +lean_inc_ref(x_1); +x_58 = l_Lean_Compiler_LCNF_LCtx_addLetDecl(x_50, 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_51); +x_60 = lean_st_ref_set(x_6, x_59); +x_61 = l_Lean_Compiler_LCNF_Code_toMono(x_25, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +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 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_62); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 1, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +return x_65; +} +else +{ +lean_dec_ref(x_1); +return x_61; +} +} +} +else +{ +uint8_t x_66; +lean_free_object(x_22); +lean_dec_ref(x_25); +lean_dec_ref(x_24); +lean_free_object(x_1); +lean_dec(x_11); +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); +x_66 = !lean_is_exclusive(x_27); +if (x_66 == 0) +{ +return x_27; +} +else +{ +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_27, 0); +lean_inc(x_67); +lean_dec(x_27); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_67); +return x_68; +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_22, 1); +x_70 = lean_ctor_get(x_22, 2); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_22); +x_71 = l_Lean_Compiler_LCNF_eraseParams___redArg(x_69, x_6); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_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_dec_ref(x_71); +x_72 = lean_st_ref_take(x_6); +x_73 = l_Lean_Compiler_LCNF_argsToMonoRedArg___closed__0; +x_74 = lean_array_get(x_73, x_69, x_21); +lean_dec_ref(x_69); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec_ref(x_74); +x_77 = lean_ctor_get(x_72, 0); +lean_inc_ref(x_77); +x_78 = lean_ctor_get(x_72, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_79 = x_72; +} else { + lean_dec_ref(x_72); + x_79 = lean_box(0); +} +x_80 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__2; +x_81 = l_Lean_Name_str___override(x_2, x_3); +x_82 = lean_box(0); +x_83 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_83, 0, x_11); +x_84 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__4; +x_85 = lean_array_push(x_84, x_83); +x_86 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_86, 0, x_81); +lean_ctor_set(x_86, 1, x_82); +lean_ctor_set(x_86, 2, x_85); +lean_ctor_set(x_1, 3, x_86); +lean_ctor_set(x_1, 2, x_80); +lean_ctor_set(x_1, 1, x_76); +lean_ctor_set(x_1, 0, x_75); +lean_inc_ref(x_1); +x_87 = l_Lean_Compiler_LCNF_LCtx_addLetDecl(x_77, x_1); +if (lean_is_scalar(x_79)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_79; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_78); +x_89 = lean_st_ref_set(x_6, x_88); +x_90 = l_Lean_Compiler_LCNF_Code_toMono(x_70, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + x_92 = x_90; +} else { + lean_dec_ref(x_90); + x_92 = lean_box(0); +} +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_91); +if (lean_is_scalar(x_92)) { + x_94 = lean_alloc_ctor(0, 1, 0); +} else { + x_94 = x_92; +} +lean_ctor_set(x_94, 0, x_93); +return x_94; +} +else +{ +lean_dec_ref(x_1); +return x_90; +} +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec_ref(x_70); +lean_dec_ref(x_69); +lean_free_object(x_1); +lean_dec(x_11); +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); +x_95 = lean_ctor_get(x_71, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + x_96 = x_71; +} else { + lean_dec_ref(x_71); + x_96 = lean_box(0); +} +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(1, 1, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_95); +return x_97; +} +} +} +else +{ +lean_object* x_98; lean_object* x_99; +lean_dec_ref(x_22); +lean_free_object(x_1); +lean_dec(x_11); +lean_dec_ref(x_3); +lean_dec(x_2); +x_98 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2; +x_99 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__1(x_98, x_4, x_5, x_6, x_7, x_8); +return x_99; +} +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; +x_100 = lean_ctor_get(x_1, 2); +x_101 = lean_ctor_get(x_1, 3); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_1); +x_102 = lean_array_get_size(x_101); +x_103 = lean_unsigned_to_nat(1u); +x_104 = lean_nat_dec_eq(x_102, x_103); +lean_dec(x_102); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec_ref(x_101); +lean_dec(x_100); +lean_dec_ref(x_3); +lean_dec(x_2); +x_105 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1; +x_106 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__1(x_105, x_4, x_5, x_6, x_7, x_8); +return x_106; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__3___closed__0; +x_108 = lean_unsigned_to_nat(0u); +x_109 = lean_array_get(x_107, x_101, x_108); +lean_dec_ref(x_101); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_109, 1); +lean_inc_ref(x_110); +x_111 = lean_ctor_get(x_109, 2); +lean_inc_ref(x_111); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + lean_ctor_release(x_109, 2); + x_112 = x_109; +} else { + lean_dec_ref(x_109); + x_112 = lean_box(0); +} +x_113 = l_Lean_Compiler_LCNF_eraseParams___redArg(x_110, x_6); +if (lean_obj_tag(x_113) == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; 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_dec_ref(x_113); +x_114 = lean_st_ref_take(x_6); +x_115 = l_Lean_Compiler_LCNF_argsToMonoRedArg___closed__0; +x_116 = lean_array_get(x_115, x_110, x_108); +lean_dec_ref(x_110); +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 = lean_ctor_get(x_114, 0); +lean_inc_ref(x_119); +x_120 = lean_ctor_get(x_114, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_121 = x_114; +} else { + lean_dec_ref(x_114); + x_121 = lean_box(0); +} +x_122 = l_Lean_Compiler_LCNF_casesTaskToMono___redArg___closed__2; +x_123 = l_Lean_Name_str___override(x_2, x_3); +x_124 = lean_box(0); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_100); +x_126 = l_Lean_Compiler_LCNF_casesStringToMono___redArg___closed__4; +x_127 = lean_array_push(x_126, x_125); +if (lean_is_scalar(x_112)) { + x_128 = lean_alloc_ctor(3, 3, 0); +} else { + x_128 = x_112; + lean_ctor_set_tag(x_128, 3); +} +lean_ctor_set(x_128, 0, x_123); +lean_ctor_set(x_128, 1, x_124); +lean_ctor_set(x_128, 2, x_127); +x_129 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_129, 0, x_117); +lean_ctor_set(x_129, 1, x_118); +lean_ctor_set(x_129, 2, x_122); +lean_ctor_set(x_129, 3, x_128); +lean_inc_ref(x_129); +x_130 = l_Lean_Compiler_LCNF_LCtx_addLetDecl(x_119, x_129); +if (lean_is_scalar(x_121)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_121; +} +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_120); +x_132 = lean_st_ref_set(x_6, x_131); +x_133 = l_Lean_Compiler_LCNF_Code_toMono(x_111, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + x_135 = x_133; +} else { + lean_dec_ref(x_133); + x_135 = lean_box(0); +} +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_129); +lean_ctor_set(x_136, 1, x_134); +if (lean_is_scalar(x_135)) { + x_137 = lean_alloc_ctor(0, 1, 0); +} else { + x_137 = x_135; +} +lean_ctor_set(x_137, 0, x_136); +return x_137; +} +else +{ +lean_dec_ref(x_129); +return x_133; +} +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_dec(x_112); +lean_dec_ref(x_111); +lean_dec_ref(x_110); +lean_dec(x_100); +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); +x_138 = lean_ctor_get(x_113, 0); +lean_inc(x_138); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + x_139 = x_113; +} else { + lean_dec_ref(x_113); + x_139 = lean_box(0); +} +if (lean_is_scalar(x_139)) { + x_140 = lean_alloc_ctor(1, 1, 0); +} else { + x_140 = x_139; +} +lean_ctor_set(x_140, 0, x_138); +return x_140; +} +} +else +{ +lean_object* x_141; lean_object* x_142; +lean_dec_ref(x_109); +lean_dec(x_100); +lean_dec_ref(x_3); +lean_dec(x_2); +x_141 = l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2; +x_142 = l_panic___at___00Lean_Compiler_LCNF_Code_toMono_spec__1(x_141, x_4, x_5, x_6, x_7, x_8); +return x_142; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_Compiler_LCNF_casesSIntToMono___redArg(x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Compiler.LCNF.casesUIntToMono", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___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_Lean_Compiler_LCNF_trivialStructToMono___closed__1; +x_2 = lean_unsigned_to_nat(2u); +x_3 = lean_unsigned_to_nat(258u); +x_4 = l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0; +x_5 = l_Lean_Compiler_LCNF_LetValue_toMono___closed__19; +x_6 = l_mkPanicMessageWithDecl(x_5, x_4, x_3, x_2, x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toBitVec", 8, 8); +return x_1; +} } static lean_object* _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__3() { _start: @@ -19848,6 +20594,22 @@ x_9 = l_Lean_Compiler_LCNF_casesArrayToMono(x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_casesSIntToMono___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_Lean_Compiler_LCNF_casesSIntToMono___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_Compiler_LCNF_casesSIntToMono___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Compiler_LCNF_casesSIntToMono(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_Compiler_LCNF_casesUIntToMono___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: { @@ -21242,6 +22004,30 @@ l_Lean_Compiler_LCNF_Code_toMono___closed__37 = _init_l_Lean_Compiler_LCNF_Code_ lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__37); l_Lean_Compiler_LCNF_Code_toMono___closed__38 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__38(); lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__38); +l_Lean_Compiler_LCNF_Code_toMono___closed__39 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__39(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__39); +l_Lean_Compiler_LCNF_Code_toMono___closed__40 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__40(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__40); +l_Lean_Compiler_LCNF_Code_toMono___closed__41 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__41(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__41); +l_Lean_Compiler_LCNF_Code_toMono___closed__42 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__42(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__42); +l_Lean_Compiler_LCNF_Code_toMono___closed__43 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__43(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__43); +l_Lean_Compiler_LCNF_Code_toMono___closed__44 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__44(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__44); +l_Lean_Compiler_LCNF_Code_toMono___closed__45 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__45(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__45); +l_Lean_Compiler_LCNF_Code_toMono___closed__46 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__46(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__46); +l_Lean_Compiler_LCNF_Code_toMono___closed__47 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__47(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__47); +l_Lean_Compiler_LCNF_Code_toMono___closed__48 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__48(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__48); +l_Lean_Compiler_LCNF_Code_toMono___closed__49 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__49(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__49); +l_Lean_Compiler_LCNF_Code_toMono___closed__50 = _init_l_Lean_Compiler_LCNF_Code_toMono___closed__50(); +lean_mark_persistent(l_Lean_Compiler_LCNF_Code_toMono___closed__50); l_Lean_Compiler_LCNF_trivialStructToMono___closed__0 = _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__0(); lean_mark_persistent(l_Lean_Compiler_LCNF_trivialStructToMono___closed__0); l_Lean_Compiler_LCNF_trivialStructToMono___closed__1 = _init_l_Lean_Compiler_LCNF_trivialStructToMono___closed__1(); @@ -21326,6 +22112,12 @@ l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__2 = _init_l_Lean_Compil lean_mark_persistent(l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__2); l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__3 = _init_l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__3(); lean_mark_persistent(l_Lean_Compiler_LCNF_casesArrayToMono___redArg___closed__3); +l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0 = _init_l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0(); +lean_mark_persistent(l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__0); +l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1 = _init_l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__1); +l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2 = _init_l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_casesSIntToMono___redArg___closed__2); l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0 = _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0(); lean_mark_persistent(l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__0); l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__1 = _init_l_Lean_Compiler_LCNF_casesUIntToMono___redArg___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Util.c b/stage0/stdlib/Lean/Compiler/LCNF/Util.c index 15e883b2f0bd..4c6c76770917 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Util.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Util.c @@ -14,11 +14,13 @@ extern "C" { #endif static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__5; +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__4; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__22; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49; static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0___redArg___closed__3; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__18; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44; @@ -27,7 +29,9 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknow LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__0; lean_object* l_Lean_Environment_header(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__14; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__13; lean_object* lean_array_push(lean_object*, lean_object*); @@ -39,6 +43,7 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnkn static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__29; extern lean_object* l_Lean_unknownIdentifierMessageTag; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58; uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l_Lean_Compiler_LCNF_isCasesApp_x3f___closed__2; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -63,6 +68,7 @@ static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknow size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getCtorArity_x3f(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__2_spec__2___redArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__37; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -77,10 +83,12 @@ LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_isRuntimeBuiltinType(lean_object*); lean_object* lean_st_ref_get(lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_getCasesInfo_x3f_spec__1___closed__0; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_isRuntimeBuiltinType___boxed(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__43; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__16; static lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0___redArg___closed__0; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__14; +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__5; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_isCompilerRelevantMData(lean_object*); LEAN_EXPORT lean_object* l_panic___at___00Lean_Compiler_LCNF_getCasesInfo_x3f_spec__0(lean_object*, lean_object*, lean_object*); @@ -126,14 +134,18 @@ LEAN_EXPORT lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkU uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__34; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__2_spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47; LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_EnvironmentHeader_moduleNames(lean_object*); lean_object* lean_panic_fn(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___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___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___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50; +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Array_contains___at___00Lean_Compiler_LCNF_isRuntimeBuiltinType_spec__0_spec__0(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__35; @@ -141,10 +153,12 @@ uint8_t l_Lean_isPrivateName(lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__17; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__21; LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Compiler_LCNF_getCasesInfo_x3f_spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(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___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__6; lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__31; +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48; lean_object* lean_array_mk(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*); @@ -187,6 +201,7 @@ uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__33; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__23; LEAN_EXPORT lean_object* l_Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46; static lean_object* l_Lean_mkUnknownIdentifierMessageCore___at___00Lean_mkUnknownIdentifierMessage___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__15; static lean_object* l_Lean_Compiler_LCNF_isCasesApp_x3f___closed__0; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00__private_Lean_Compiler_LCNF_Util_0__Lean_Compiler_LCNF_getCasesOnInductiveVal_x3f_spec__0_spec__0_spec__0___redArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2296,7 +2311,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__12( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float", 5, 5); +x_1 = lean_mk_string_unchecked("Int8", 4, 4); return x_1; } } @@ -2313,7 +2328,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__14( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Float32", 7, 7); +x_1 = lean_mk_string_unchecked("Int16", 5, 5); return x_1; } } @@ -2330,7 +2345,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__16( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Thunk", 5, 5); +x_1 = lean_mk_string_unchecked("Int32", 5, 5); return x_1; } } @@ -2347,7 +2362,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__18( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Task", 4, 4); +x_1 = lean_mk_string_unchecked("Int64", 5, 5); return x_1; } } @@ -2364,7 +2379,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__20( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Array", 5, 5); +x_1 = lean_mk_string_unchecked("ISize", 5, 5); return x_1; } } @@ -2381,7 +2396,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__22( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("ByteArray", 9, 9); +x_1 = lean_mk_string_unchecked("Float", 5, 5); return x_1; } } @@ -2398,7 +2413,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__24( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("FloatArray", 10, 10); +x_1 = lean_mk_string_unchecked("Float32", 7, 7); return x_1; } } @@ -2415,7 +2430,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__26( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Nat", 3, 3); +x_1 = lean_mk_string_unchecked("Thunk", 5, 5); return x_1; } } @@ -2432,7 +2447,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__28( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Int", 3, 3); +x_1 = lean_mk_string_unchecked("Task", 4, 4); return x_1; } } @@ -2448,158 +2463,293 @@ return x_2; static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__30() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Array", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__30; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__32() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ByteArray", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__32; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("FloatArray", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__34; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Nat", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__36; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__38() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Int", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__38; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__40() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(15u); +x_1 = lean_unsigned_to_nat(20u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__31() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__1; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__30; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__40; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__32() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__3; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__31; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__41; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__33() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__5; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__32; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__42; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__34() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__7; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__33; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__43; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__35() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__9; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__34; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__36() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__11; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__35; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__37() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__13; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__36; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__38() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__15; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__37; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__39() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__17; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__38; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__40() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__19; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__39; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__41() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__21; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__40; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__42() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__23; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__41; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__43() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__25; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__42; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__27; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__43; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53; x_3 = lean_array_push(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45() { +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__29; -x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54; +x_3 = lean_array_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__31; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55; +x_3 = lean_array_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__33; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56; +x_3 = lean_array_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__35; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57; +x_3 = lean_array_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__37; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58; +x_3 = lean_array_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__39; +x_2 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59; x_3 = lean_array_push(x_2, x_1); return x_3; } @@ -2608,7 +2758,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes() { _start: { lean_object* x_1; -x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45; +x_1 = l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60; return x_1; } } @@ -2901,6 +3051,36 @@ l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44 = _init_l_Lean_Compiler_LC lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__44); l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45(); lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__45); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__46); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__47); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__48); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__49); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__50); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__51); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__52); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__53); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__54); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__55); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__56); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__57); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__58); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__59); +l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60 = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60(); +lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__60); l_Lean_Compiler_LCNF_builtinRuntimeTypes = _init_l_Lean_Compiler_LCNF_builtinRuntimeTypes(); lean_mark_persistent(l_Lean_Compiler_LCNF_builtinRuntimeTypes); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index 328f3d7a2f75..83a13b5ae377 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -11174,21 +11174,21 @@ goto block_40; } block_26: { -if (lean_obj_tag(x_16) == 0) +if (lean_obj_tag(x_17) == 0) { lean_object* x_21; -x_21 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux(x_20, x_11, x_10, x_18, x_17, x_14, x_15, x_13, x_12); +x_21 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux(x_20, x_12, x_10, x_15, x_14, x_19, x_16, x_18, x_11); return x_21; } else { lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_16, 0); +x_22 = lean_ctor_get(x_17, 0); lean_inc(x_22); -lean_dec_ref(x_16); +lean_dec_ref(x_17); x_23 = l_Lean_Elab_Term_elabLeadingParserMacro___lam__0___closed__3; x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -x_25 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux(x_20, x_11, x_24, x_18, x_17, x_14, x_15, x_13, x_12); +x_25 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux(x_20, x_12, x_24, x_15, x_14, x_19, x_16, x_18, x_11); return x_25; } } @@ -11202,15 +11202,15 @@ if (lean_obj_tag(x_27) == 0) { lean_object* x_38; x_38 = l_Lean_Elab_Term_elabLeadingParserMacro___lam__0___closed__6; -x_11 = x_37; -x_12 = x_34; -x_13 = x_33; -x_14 = x_31; -x_15 = x_32; -x_16 = x_28; -x_17 = x_30; -x_18 = x_29; -x_19 = lean_box(0); +x_11 = x_34; +x_12 = x_37; +x_13 = lean_box(0); +x_14 = x_30; +x_15 = x_29; +x_16 = x_32; +x_17 = x_28; +x_18 = x_33; +x_19 = x_31; x_20 = x_38; goto block_26; } @@ -11220,15 +11220,15 @@ lean_object* x_39; x_39 = lean_ctor_get(x_27, 0); lean_inc(x_39); lean_dec_ref(x_27); -x_11 = x_37; -x_12 = x_34; -x_13 = x_33; -x_14 = x_31; -x_15 = x_32; -x_16 = x_28; -x_17 = x_30; -x_18 = x_29; -x_19 = lean_box(0); +x_11 = x_34; +x_12 = x_37; +x_13 = lean_box(0); +x_14 = x_30; +x_15 = x_29; +x_16 = x_32; +x_17 = x_28; +x_18 = x_33; +x_19 = x_31; x_20 = x_39; goto block_26; } @@ -11763,7 +11763,7 @@ if (lean_obj_tag(x_20) == 0) { lean_object* x_22; lean_object* x_23; x_22 = l_Lean_Elab_Term_elabTrailingParserMacro___lam__0___closed__3; -x_23 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux(x_21, x_22, x_12, x_17, x_14, x_15, x_19, x_13, x_18); +x_23 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux(x_21, x_22, x_17, x_14, x_18, x_13, x_16, x_12, x_15); return x_23; } else @@ -11772,7 +11772,7 @@ lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); lean_dec_ref(x_20); -x_25 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux(x_21, x_24, x_12, x_17, x_14, x_15, x_19, x_13, x_18); +x_25 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux(x_21, x_24, x_17, x_14, x_18, x_13, x_16, x_12, x_15); return x_25; } } @@ -11786,14 +11786,14 @@ if (lean_obj_tag(x_27) == 0) { lean_object* x_38; x_38 = l_Lean_Elab_Term_elabLeadingParserMacro___lam__0___closed__6; -x_12 = x_37; -x_13 = x_33; -x_14 = x_30; -x_15 = x_31; -x_16 = lean_box(0); -x_17 = x_29; -x_18 = x_34; -x_19 = x_32; +x_12 = x_33; +x_13 = x_31; +x_14 = x_29; +x_15 = x_34; +x_16 = x_32; +x_17 = x_37; +x_18 = x_30; +x_19 = lean_box(0); x_20 = x_28; x_21 = x_38; goto block_26; @@ -11804,14 +11804,14 @@ lean_object* x_39; x_39 = lean_ctor_get(x_27, 0); lean_inc(x_39); lean_dec_ref(x_27); -x_12 = x_37; -x_13 = x_33; -x_14 = x_30; -x_15 = x_31; -x_16 = lean_box(0); -x_17 = x_29; -x_18 = x_34; -x_19 = x_32; +x_12 = x_33; +x_13 = x_31; +x_14 = x_29; +x_15 = x_34; +x_16 = x_32; +x_17 = x_37; +x_18 = x_30; +x_19 = lean_box(0); x_20 = x_28; x_21 = x_39; goto block_26; @@ -19487,16 +19487,16 @@ if (x_178 == 0) uint8_t x_179; x_179 = l_Lean_Exception_isRuntime(x_177); lean_dec(x_177); -x_21 = x_176; -x_22 = lean_box(0); +x_21 = lean_box(0); +x_22 = x_176; x_23 = x_179; goto block_26; } else { lean_dec(x_177); -x_21 = x_176; -x_22 = lean_box(0); +x_21 = lean_box(0); +x_22 = x_176; x_23 = x_178; goto block_26; } @@ -19624,16 +19624,16 @@ if (x_210 == 0) uint8_t x_211; x_211 = l_Lean_Exception_isRuntime(x_209); lean_dec(x_209); -x_27 = lean_box(0); -x_28 = x_208; +x_27 = x_208; +x_28 = lean_box(0); x_29 = x_211; goto block_32; } else { lean_dec(x_209); -x_27 = lean_box(0); -x_28 = x_208; +x_27 = x_208; +x_28 = lean_box(0); x_29 = x_210; goto block_32; } @@ -19898,16 +19898,16 @@ if (x_274 == 0) uint8_t x_275; x_275 = l_Lean_Exception_isRuntime(x_273); lean_dec(x_273); -x_39 = lean_box(0); -x_40 = x_272; +x_39 = x_272; +x_40 = lean_box(0); x_41 = x_275; goto block_44; } else { lean_dec(x_273); -x_39 = lean_box(0); -x_40 = x_272; +x_39 = x_272; +x_40 = lean_box(0); x_41 = x_274; goto block_44; } @@ -20034,16 +20034,16 @@ if (x_306 == 0) uint8_t x_307; x_307 = l_Lean_Exception_isRuntime(x_305); lean_dec(x_305); -x_45 = x_304; -x_46 = lean_box(0); +x_45 = lean_box(0); +x_46 = x_304; x_47 = x_307; goto block_50; } else { lean_dec(x_305); -x_45 = x_304; -x_46 = lean_box(0); +x_45 = lean_box(0); +x_46 = x_304; x_47 = x_306; goto block_50; } @@ -20362,7 +20362,7 @@ return x_15; if (x_23 == 0) { lean_object* x_24; lean_object* x_25; -lean_dec_ref(x_21); +lean_dec_ref(x_22); x_24 = lean_box(0); x_25 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_25, 0, x_24); @@ -20370,7 +20370,7 @@ return x_25; } else { -return x_21; +return x_22; } } block_32: @@ -20378,7 +20378,7 @@ return x_21; if (x_29 == 0) { lean_object* x_30; lean_object* x_31; -lean_dec_ref(x_28); +lean_dec_ref(x_27); x_30 = lean_box(0); x_31 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_31, 0, x_30); @@ -20386,7 +20386,7 @@ return x_31; } else { -return x_28; +return x_27; } } block_38: @@ -20410,7 +20410,7 @@ return x_33; if (x_41 == 0) { lean_object* x_42; lean_object* x_43; -lean_dec_ref(x_40); +lean_dec_ref(x_39); x_42 = lean_box(0); x_43 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_43, 0, x_42); @@ -20418,7 +20418,7 @@ return x_43; } else { -return x_40; +return x_39; } } block_50: @@ -20426,7 +20426,7 @@ return x_40; if (x_47 == 0) { lean_object* x_48; lean_object* x_49; -lean_dec_ref(x_45); +lean_dec_ref(x_46); x_48 = lean_box(0); x_49 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_49, 0, x_48); @@ -20434,7 +20434,7 @@ return x_49; } else { -return x_45; +return x_46; } } } @@ -21998,13 +21998,13 @@ if (lean_obj_tag(x_33) == 0) { lean_object* x_60; x_60 = l_Lean_Elab_Term_expandTypeAscription___closed__0; -x_5 = x_41; -x_6 = x_59; -x_7 = x_57; -x_8 = x_58; -x_9 = x_46; -x_10 = x_55; -x_11 = x_40; +x_5 = x_58; +x_6 = x_57; +x_7 = x_59; +x_8 = x_41; +x_9 = x_40; +x_10 = x_46; +x_11 = x_55; x_12 = x_60; goto block_19; } @@ -22016,13 +22016,13 @@ lean_inc(x_61); lean_dec_ref(x_33); x_62 = l_Lean_Elab_Term_expandTypeAscription___closed__0; x_63 = lean_array_push(x_62, x_61); -x_5 = x_41; -x_6 = x_59; -x_7 = x_57; -x_8 = x_58; -x_9 = x_46; -x_10 = x_55; -x_11 = x_40; +x_5 = x_58; +x_6 = x_57; +x_7 = x_59; +x_8 = x_41; +x_9 = x_40; +x_10 = x_46; +x_11 = x_55; x_12 = x_63; goto block_19; } @@ -22059,22 +22059,22 @@ return x_67; block_19: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_13 = l_Array_append___redArg(x_6, x_12); +x_13 = l_Array_append___redArg(x_7, x_12); lean_dec_ref(x_12); -lean_inc(x_9); +lean_inc(x_10); x_14 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_14, 0, x_9); -lean_ctor_set(x_14, 1, x_8); +lean_ctor_set(x_14, 0, x_10); +lean_ctor_set(x_14, 1, x_5); lean_ctor_set(x_14, 2, x_13); x_15 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__35; -lean_inc(x_9); +lean_inc(x_10); x_16 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 0, x_10); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Syntax_node5(x_9, x_4, x_10, x_5, x_7, x_14, x_16); +x_17 = l_Lean_Syntax_node5(x_10, x_4, x_11, x_8, x_6, x_14, x_16); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_11); +lean_ctor_set(x_18, 1, x_9); return x_18; } } @@ -25269,20 +25269,20 @@ if (x_42 == 0) if (x_2 == 0) { lean_object* x_46; -lean_dec_ref(x_41); +lean_dec_ref(x_43); lean_dec(x_20); lean_dec_ref(x_19); lean_dec_ref(x_17); lean_dec_ref(x_16); lean_dec(x_3); -x_46 = l_Lean_Meta_mkEqRec(x_43, x_44, x_15, x_21, x_22, x_23, x_24); +x_46 = l_Lean_Meta_mkEqRec(x_44, x_41, x_15, x_21, x_22, x_23, x_24); return x_46; } else { lean_dec_ref(x_44); -lean_dec_ref(x_43); -x_34 = x_41; +lean_dec_ref(x_41); +x_34 = x_43; x_35 = lean_box(0); goto block_40; } @@ -25290,13 +25290,13 @@ goto block_40; else { lean_object* x_47; -lean_dec_ref(x_41); +lean_dec_ref(x_43); lean_dec(x_20); lean_dec_ref(x_19); lean_dec_ref(x_17); lean_dec_ref(x_16); lean_dec(x_3); -x_47 = l_Lean_Meta_mkEqRec(x_43, x_44, x_15, x_21, x_22, x_23, x_24); +x_47 = l_Lean_Meta_mkEqRec(x_44, x_41, x_15, x_21, x_22, x_23, x_24); return x_47; } } @@ -25348,10 +25348,10 @@ if (lean_obj_tag(x_50) == 0) uint8_t x_59; x_59 = lean_unbox(x_55); lean_dec(x_55); -x_41 = x_58; +x_41 = x_49; x_42 = x_59; -x_43 = x_53; -x_44 = x_49; +x_43 = x_58; +x_44 = x_53; x_45 = lean_box(0); goto block_48; } @@ -25363,10 +25363,10 @@ if (x_2 == 0) uint8_t x_60; x_60 = lean_unbox(x_55); lean_dec(x_55); -x_41 = x_58; +x_41 = x_49; x_42 = x_60; -x_43 = x_53; -x_44 = x_49; +x_43 = x_58; +x_44 = x_53; x_45 = lean_box(0); goto block_48; } diff --git a/stage0/stdlib/Lean/Elab/Calc.c b/stage0/stdlib/Lean/Elab/Calc.c index f5f54ba60992..d406ab159395 100644 --- a/stage0/stdlib/Lean/Elab/Calc.c +++ b/stage0/stdlib/Lean/Elab/Calc.c @@ -10660,7 +10660,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_Term_throwCalcFailure_spec__2_spec__2(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; 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; uint8_t x_51; lean_object* 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; lean_object* x_80; uint8_t 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; 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_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; 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; 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; 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) @@ -10695,15 +10695,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_15); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_16); -lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_10); +lean_ctor_set(x_27, 0, x_11); +lean_ctor_set(x_27, 1, x_14); +lean_ctor_set(x_27, 2, x_10); +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_12); -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_16); 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); @@ -10740,15 +10740,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_15); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_14); -lean_ctor_set(x_43, 1, x_16); -lean_ctor_set(x_43, 2, x_15); -lean_ctor_set(x_43, 3, x_10); +lean_ctor_set(x_43, 0, x_11); +lean_ctor_set(x_43, 1, x_14); +lean_ctor_set(x_43, 2, x_10); +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_12); -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_16); 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); @@ -10778,25 +10778,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_52); -x_62 = l_Lean_FileMap_toPosition(x_52, x_56); -lean_dec(x_56); -x_63 = l_Lean_FileMap_toPosition(x_52, x_57); +lean_inc_ref(x_53); +x_62 = l_Lean_FileMap_toPosition(x_53, x_54); +lean_dec(x_54); +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___private_Lean_Elab_Calc_0__Lean_Elab_Term_annotateFirstHoleWithType_go___closed__11; -if (x_51 == 0) +if (x_55 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_65; -x_11 = x_61; -x_12 = x_53; -x_13 = x_55; -x_14 = x_54; -x_15 = x_64; -x_16 = x_62; +x_10 = x_64; +x_11 = x_51; +x_12 = x_65; +x_13 = x_52; +x_14 = x_62; +x_15 = x_61; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -10813,7 +10813,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); @@ -10822,13 +10822,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_65; -x_11 = x_61; -x_12 = x_53; -x_13 = x_55; -x_14 = x_54; -x_15 = x_64; -x_16 = x_62; +x_10 = x_64; +x_11 = x_51; +x_12 = x_65; +x_13 = x_52; +x_14 = x_62; +x_15 = x_61; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -10842,24 +10842,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_52); -x_69 = l_Lean_FileMap_toPosition(x_52, x_56); -lean_dec(x_56); -x_70 = l_Lean_FileMap_toPosition(x_52, x_57); +lean_inc_ref(x_53); +x_69 = l_Lean_FileMap_toPosition(x_53, x_54); +lean_dec(x_54); +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___private_Lean_Elab_Calc_0__Lean_Elab_Term_annotateFirstHoleWithType_go___closed__11; -if (x_51 == 0) +if (x_55 == 0) { lean_dec_ref(x_50); -x_10 = x_72; -x_11 = x_68; -x_12 = x_53; -x_13 = x_55; -x_14 = x_54; -x_15 = x_71; -x_16 = x_69; +x_10 = x_71; +x_11 = x_51; +x_12 = x_72; +x_13 = x_52; +x_14 = x_69; +x_15 = x_68; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -10876,7 +10876,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); @@ -10885,13 +10885,13 @@ return x_75; } else { -x_10 = x_72; -x_11 = x_68; -x_12 = x_53; -x_13 = x_55; -x_14 = x_54; -x_15 = x_71; -x_16 = x_69; +x_10 = x_71; +x_11 = x_51; +x_12 = x_72; +x_13 = x_52; +x_14 = x_69; +x_15 = x_68; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -10903,18 +10903,18 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_79, x_81); -lean_dec(x_79); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_78, x_81); +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_80; -x_53 = x_81; -x_54 = x_83; +x_51 = x_79; +x_52 = x_81; +x_53 = x_80; +x_54 = x_84; x_55 = x_82; -x_56 = x_84; +x_56 = x_83; x_57 = x_84; goto block_76; } @@ -10925,12 +10925,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_80; -x_53 = x_81; -x_54 = x_83; +x_51 = x_79; +x_52 = x_81; +x_53 = x_80; +x_54 = x_84; x_55 = x_82; -x_56 = x_84; +x_56 = x_83; x_57 = x_86; goto block_76; } @@ -10938,20 +10938,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_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_95; -x_80 = x_90; -x_81 = x_91; -x_82 = x_94; -x_83 = x_92; +x_78 = x_95; +x_79 = x_89; +x_80 = x_91; +x_81 = x_90; +x_82 = x_93; +x_83 = x_94; x_84 = x_97; goto block_87; } @@ -10962,12 +10962,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_95; -x_80 = x_90; -x_81 = x_91; -x_82 = x_94; -x_83 = x_92; +x_78 = x_95; +x_79 = x_89; +x_80 = x_91; +x_81 = x_90; +x_82 = x_93; +x_83 = x_94; x_84 = x_98; goto block_87; } @@ -10976,23 +10976,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_104; -x_93 = x_105; +x_88 = x_105; +x_89 = x_101; +x_90 = x_106; +x_91 = x_102; +x_92 = x_103; +x_93 = x_104; 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_104; -x_93 = x_105; +x_88 = x_105; +x_89 = x_101; +x_90 = x_106; +x_91 = x_102; +x_92 = x_103; +x_93 = x_104; x_94 = x_100; goto block_99; } @@ -11017,13 +11017,13 @@ 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_ref(x_111); -x_101 = x_117; -x_102 = x_114; -x_103 = x_111; -x_104 = x_110; -x_105 = x_113; +lean_inc_ref(x_110); +x_101 = x_110; +x_102 = x_111; +x_103 = x_113; +x_104 = x_114; +x_105 = x_117; x_106 = x_109; x_107 = x_119; goto block_108; @@ -11034,13 +11034,13 @@ lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_Term_throwCalcFailure_spec__2_spec__2___closed__0; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_Term_elabCalcSteps_spec__0_spec__0_spec__0_spec__0(x_112, x_120); lean_inc(x_113); -lean_inc_ref(x_110); lean_inc_ref(x_111); -x_101 = x_117; -x_102 = x_114; -x_103 = x_111; -x_104 = x_110; -x_105 = x_113; +lean_inc_ref(x_110); +x_101 = x_110; +x_102 = x_111; +x_103 = x_113; +x_104 = x_114; +x_105 = x_117; x_106 = x_109; x_107 = x_121; goto block_108; diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index 9c738532c1ea..0e0136feaad8 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -20055,7 +20055,7 @@ return x_152; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; 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_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_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; uint8_t x_259; lean_object* x_260; 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; uint8_t x_302; lean_object* x_303; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; uint8_t x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; uint8_t x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; uint8_t x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; uint8_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; uint8_t x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; uint8_t 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_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; uint8_t x_528; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; uint8_t x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; +lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; 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_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; 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; uint8_t x_301; lean_object* x_302; lean_object* x_303; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; uint8_t x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; uint8_t x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_362; lean_object* x_363; lean_object* x_364; uint8_t 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_391; lean_object* x_392; uint8_t x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_416; lean_object* x_417; uint8_t 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_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; uint8_t 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_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; uint8_t x_528; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; uint8_t x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; x_153 = lean_unsigned_to_nat(1u); x_154 = l_Lean_Syntax_getArg(x_1, x_153); x_155 = l_Lean_Elab_Command_elabInitialize___closed__40; @@ -20074,7 +20074,7 @@ return x_571; } else { -lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; uint8_t x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; 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; uint8_t x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; uint8_t x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_749; lean_object* x_750; lean_object* x_751; 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_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_773; 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_object* x_785; lean_object* x_786; 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; uint8_t 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_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; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; uint8_t x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; uint8_t x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_969; uint8_t x_970; +lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; uint8_t x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; uint8_t x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; 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; uint8_t x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; uint8_t x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_749; lean_object* x_750; lean_object* x_751; 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_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_773; 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_object* x_785; lean_object* x_786; 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; uint8_t 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_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; uint8_t x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; uint8_t x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_888; lean_object* x_889; lean_object* x_890; uint8_t x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_969; uint8_t x_970; x_572 = lean_unsigned_to_nat(2u); x_969 = l_Lean_Syntax_getArg(x_1, x_572); x_970 = l_Lean_Syntax_isNone(x_969); @@ -20151,7 +20151,7 @@ goto block_968; block_633: { lean_object* x_596; -x_596 = l_Lean_Elab_Command_getScope___redArg(x_578); +x_596 = l_Lean_Elab_Command_getScope___redArg(x_589); if (lean_obj_tag(x_596) == 0) { lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; @@ -20161,35 +20161,35 @@ lean_dec_ref(x_596); x_598 = lean_ctor_get(x_597, 2); lean_inc(x_598); lean_dec(x_597); -lean_inc_ref(x_587); -x_599 = l_Array_append___redArg(x_587, x_595); +lean_inc_ref(x_573); +x_599 = l_Array_append___redArg(x_573, x_595); lean_dec_ref(x_595); -lean_inc(x_589); -lean_inc(x_577); +lean_inc(x_581); +lean_inc(x_584); x_600 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_600, 0, x_577); -lean_ctor_set(x_600, 1, x_589); +lean_ctor_set(x_600, 0, x_584); +lean_ctor_set(x_600, 1, x_581); lean_ctor_set(x_600, 2, x_599); -lean_inc_n(x_574, 3); -lean_inc(x_577); -x_601 = l_Lean_Syntax_node7(x_577, x_13, x_584, x_586, x_590, x_574, x_600, x_574, x_574); +lean_inc_n(x_594, 3); +lean_inc(x_584); +x_601 = l_Lean_Syntax_node7(x_584, x_13, x_578, x_575, x_586, x_594, x_600, x_594, x_594); x_602 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__5; x_603 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__6; -lean_inc(x_577); +lean_inc(x_584); x_604 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_604, 0, x_577); +lean_ctor_set(x_604, 0, x_584); lean_ctor_set(x_604, 1, x_602); x_605 = l_Lean_Elab_Command_elabInitialize___closed__6; x_606 = l_Lean_Elab_Command_expandMutualElement___closed__0; x_607 = lean_box(2); -lean_inc(x_589); +lean_inc(x_581); x_608 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_608, 0, x_607); -lean_ctor_set(x_608, 1, x_589); +lean_ctor_set(x_608, 1, x_581); lean_ctor_set(x_608, 2, x_606); x_609 = l_Lean_Elab_addDeclarationRangesForBuiltin___at___00Lean_Elab_Command_elabAxiom_spec__26___closed__2; -lean_inc(x_573); -x_610 = lean_array_push(x_609, x_573); +lean_inc(x_592); +x_610 = lean_array_push(x_609, x_592); x_611 = lean_array_push(x_610, x_608); x_612 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_612, 0, x_607); @@ -20197,111 +20197,111 @@ lean_ctor_set(x_612, 1, x_605); lean_ctor_set(x_612, 2, x_611); x_613 = l_Lean_Elab_Command_elabInitialize___closed__57; x_614 = l_Lean_Elab_Command_elabInitialize___closed__12; -lean_inc_ref(x_579); -x_615 = l_Lean_Name_mkStr4(x_5, x_6, x_579, x_614); +lean_inc_ref(x_582); +x_615 = l_Lean_Name_mkStr4(x_5, x_6, x_582, x_614); x_616 = l_Lean_Elab_Command_elabInitialize___closed__13; -lean_inc(x_577); +lean_inc(x_584); x_617 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_617, 0, x_577); +lean_ctor_set(x_617, 0, x_584); lean_ctor_set(x_617, 1, x_616); -lean_inc(x_591); +lean_inc(x_585); lean_inc(x_615); -lean_inc(x_577); -x_618 = l_Lean_Syntax_node2(x_577, x_615, x_617, x_591); -lean_inc(x_574); -lean_inc(x_577); -x_619 = l_Lean_Syntax_node2(x_577, x_613, x_574, x_618); -lean_inc(x_577); -x_620 = l_Lean_Syntax_node4(x_577, x_603, x_604, x_612, x_619, x_574); -lean_inc(x_575); -x_621 = l_Lean_Syntax_node2(x_577, x_575, x_601, x_620); -x_622 = l_Lean_TSyntax_getId(x_573); -lean_dec(x_573); +lean_inc(x_584); +x_618 = l_Lean_Syntax_node2(x_584, x_615, x_617, x_585); +lean_inc(x_594); +lean_inc(x_584); +x_619 = l_Lean_Syntax_node2(x_584, x_613, x_594, x_618); +lean_inc(x_584); +x_620 = l_Lean_Syntax_node4(x_584, x_603, x_604, x_612, x_619, x_594); +lean_inc(x_587); +x_621 = l_Lean_Syntax_node2(x_584, x_587, x_601, x_620); +x_622 = l_Lean_TSyntax_getId(x_592); +lean_dec(x_592); x_623 = l_Lean_Name_append(x_598, x_622); -if (lean_obj_tag(x_593) == 0) +if (lean_obj_tag(x_579) == 0) { -x_546 = x_615; -x_547 = x_575; -x_548 = x_576; -x_549 = x_579; +x_546 = x_573; +x_547 = x_616; +x_548 = x_615; +x_549 = x_577; x_550 = x_580; -x_551 = x_582; -x_552 = x_581; -x_553 = x_616; -x_554 = x_587; -x_555 = x_588; -x_556 = x_589; -x_557 = x_591; -x_558 = x_621; -x_559 = x_592; -x_560 = x_605; -x_561 = x_594; +x_551 = x_581; +x_552 = x_582; +x_553 = x_583; +x_554 = x_585; +x_555 = x_587; +x_556 = x_588; +x_557 = x_621; +x_558 = x_591; +x_559 = x_590; +x_560 = x_593; +x_561 = x_605; x_562 = x_623; -x_563 = x_585; -x_564 = x_578; +x_563 = x_574; +x_564 = x_589; x_565 = lean_box(0); goto block_570; } else { lean_object* x_624; lean_object* x_625; uint8_t x_626; -x_624 = lean_ctor_get(x_593, 0); +x_624 = lean_ctor_get(x_579, 0); lean_inc(x_624); -lean_dec_ref(x_593); +lean_dec_ref(x_579); x_625 = l_Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4___closed__0; x_626 = l_Lean_Syntax_isOfKind(x_624, x_625); if (x_626 == 0) { -x_546 = x_615; -x_547 = x_575; -x_548 = x_576; -x_549 = x_579; +x_546 = x_573; +x_547 = x_616; +x_548 = x_615; +x_549 = x_577; x_550 = x_580; -x_551 = x_582; -x_552 = x_581; -x_553 = x_616; -x_554 = x_587; -x_555 = x_588; -x_556 = x_589; -x_557 = x_591; -x_558 = x_621; -x_559 = x_592; -x_560 = x_605; -x_561 = x_594; +x_551 = x_581; +x_552 = x_582; +x_553 = x_583; +x_554 = x_585; +x_555 = x_587; +x_556 = x_588; +x_557 = x_621; +x_558 = x_591; +x_559 = x_590; +x_560 = x_593; +x_561 = x_605; x_562 = x_623; -x_563 = x_585; -x_564 = x_578; +x_563 = x_574; +x_564 = x_589; x_565 = lean_box(0); goto block_570; } else { lean_object* x_627; lean_object* x_628; lean_object* x_629; -x_627 = lean_st_ref_get(x_578); +x_627 = lean_st_ref_get(x_589); x_628 = lean_ctor_get(x_627, 0); lean_inc_ref(x_628); lean_dec_ref(x_627); x_629 = l_Lean_mkPrivateName(x_628, x_623); lean_dec_ref(x_628); -x_546 = x_615; -x_547 = x_575; -x_548 = x_576; -x_549 = x_579; +x_546 = x_573; +x_547 = x_616; +x_548 = x_615; +x_549 = x_577; x_550 = x_580; -x_551 = x_582; -x_552 = x_581; -x_553 = x_616; -x_554 = x_587; -x_555 = x_588; -x_556 = x_589; -x_557 = x_591; -x_558 = x_621; -x_559 = x_592; -x_560 = x_605; -x_561 = x_594; +x_551 = x_581; +x_552 = x_582; +x_553 = x_583; +x_554 = x_585; +x_555 = x_587; +x_556 = x_588; +x_557 = x_621; +x_558 = x_591; +x_559 = x_590; +x_560 = x_593; +x_561 = x_605; x_562 = x_629; -x_563 = x_585; -x_564 = x_578; +x_563 = x_574; +x_564 = x_589; x_565 = lean_box(0); goto block_570; } @@ -20315,22 +20315,22 @@ lean_dec(x_594); lean_dec(x_593); lean_dec(x_592); lean_dec(x_591); -lean_dec(x_590); +lean_dec_ref(x_590); lean_dec(x_589); -lean_dec_ref(x_588); -lean_dec_ref(x_587); +lean_dec(x_588); +lean_dec(x_587); lean_dec(x_586); -lean_dec_ref(x_585); +lean_dec(x_585); lean_dec(x_584); -lean_dec(x_582); +lean_dec(x_583); +lean_dec_ref(x_582); lean_dec(x_581); -lean_dec_ref(x_579); +lean_dec(x_579); lean_dec(x_578); lean_dec(x_577); -lean_dec(x_576); lean_dec(x_575); -lean_dec(x_574); -lean_dec(x_573); +lean_dec_ref(x_574); +lean_dec_ref(x_573); x_630 = !lean_is_exclusive(x_596); if (x_630 == 0) { @@ -20351,79 +20351,79 @@ return x_632; block_666: { lean_object* x_656; lean_object* x_657; -lean_inc_ref(x_648); -x_656 = l_Array_append___redArg(x_648, x_655); +lean_inc_ref(x_634); +x_656 = l_Array_append___redArg(x_634, x_655); lean_dec_ref(x_655); -lean_inc(x_650); -lean_inc(x_637); +lean_inc(x_640); +lean_inc(x_644); x_657 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_657, 0, x_637); -lean_ctor_set(x_657, 1, x_650); +lean_ctor_set(x_657, 0, x_644); +lean_ctor_set(x_657, 1, x_640); lean_ctor_set(x_657, 2, x_656); -if (lean_obj_tag(x_654) == 0) +if (lean_obj_tag(x_650) == 0) { lean_object* x_658; x_658 = l_Lean_Elab_Command_elabInitialize___closed__41; x_573 = x_634; x_574 = x_635; x_575 = x_636; -x_576 = x_638; -x_577 = x_637; +x_576 = lean_box(0); +x_577 = x_638; x_578 = x_639; -x_579 = x_640; -x_580 = x_641; -x_581 = x_643; -x_582 = x_644; -x_583 = lean_box(0); -x_584 = x_642; -x_585 = x_647; -x_586 = x_646; -x_587 = x_648; +x_579 = x_641; +x_580 = x_642; +x_581 = x_640; +x_582 = x_643; +x_583 = x_645; +x_584 = x_644; +x_585 = x_646; +x_586 = x_657; +x_587 = x_647; x_588 = x_649; -x_589 = x_650; -x_590 = x_657; -x_591 = x_651; +x_589 = x_648; +x_590 = x_651; +x_591 = x_650; x_592 = x_652; -x_593 = x_653; -x_594 = x_654; +x_593 = x_654; +x_594 = x_653; x_595 = x_658; goto block_633; } else { lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; -x_659 = lean_ctor_get(x_654, 0); +x_659 = lean_ctor_get(x_650, 0); x_660 = l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__12_spec__12___closed__16; x_661 = l_Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4___closed__7; x_662 = l_Lean_SourceInfo_fromRef(x_659, x_156); x_663 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_663, 0, x_662); lean_ctor_set(x_663, 1, x_660); -lean_inc(x_637); -x_664 = l_Lean_Syntax_node1(x_637, x_661, x_663); +lean_inc(x_644); +x_664 = l_Lean_Syntax_node1(x_644, x_661, x_663); x_665 = l_Array_mkArray1___redArg(x_664); x_573 = x_634; x_574 = x_635; x_575 = x_636; -x_576 = x_638; -x_577 = x_637; +x_576 = lean_box(0); +x_577 = x_638; x_578 = x_639; -x_579 = x_640; -x_580 = x_641; -x_581 = x_643; -x_582 = x_644; -x_583 = lean_box(0); -x_584 = x_642; -x_585 = x_647; -x_586 = x_646; -x_587 = x_648; +x_579 = x_641; +x_580 = x_642; +x_581 = x_640; +x_582 = x_643; +x_583 = x_645; +x_584 = x_644; +x_585 = x_646; +x_586 = x_657; +x_587 = x_647; x_588 = x_649; -x_589 = x_650; -x_590 = x_657; -x_591 = x_651; +x_589 = x_648; +x_590 = x_651; +x_591 = x_650; x_592 = x_652; -x_593 = x_653; -x_594 = x_654; +x_593 = x_654; +x_594 = x_653; x_595 = x_665; goto block_633; } @@ -20431,80 +20431,80 @@ goto block_633; block_701: { lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; -x_691 = l_Array_append___redArg(x_667, x_690); +x_691 = l_Array_append___redArg(x_673, x_690); lean_dec_ref(x_690); -lean_inc(x_684); -lean_inc(x_671); +lean_inc(x_675); +lean_inc(x_680); x_692 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_692, 0, x_671); -lean_ctor_set(x_692, 1, x_684); +lean_ctor_set(x_692, 0, x_680); +lean_ctor_set(x_692, 1, x_675); lean_ctor_set(x_692, 2, x_691); x_693 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Command_elabAttr_spec__0___closed__4; -lean_inc(x_671); +lean_inc(x_680); x_694 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_694, 0, x_671); +lean_ctor_set(x_694, 0, x_680); lean_ctor_set(x_694, 1, x_693); -lean_inc(x_671); -x_695 = l_Lean_Syntax_node3(x_671, x_673, x_688, x_692, x_694); -lean_inc(x_684); -lean_inc(x_671); -x_696 = l_Lean_Syntax_node1(x_671, x_684, x_695); -if (lean_obj_tag(x_687) == 0) +lean_inc(x_680); +x_695 = l_Lean_Syntax_node3(x_680, x_671, x_669, x_692, x_694); +lean_inc(x_675); +lean_inc(x_680); +x_696 = l_Lean_Syntax_node1(x_680, x_675, x_695); +if (lean_obj_tag(x_676) == 0) { lean_object* x_697; x_697 = l_Lean_Elab_Command_elabInitialize___closed__41; -x_634 = x_668; -x_635 = x_669; -x_636 = x_670; -x_637 = x_671; +x_634 = x_667; +x_635 = x_668; +x_636 = x_696; +x_637 = lean_box(0); x_638 = x_672; x_639 = x_674; x_640 = x_675; x_641 = x_676; -x_642 = x_680; -x_643 = x_679; -x_644 = x_678; -x_645 = lean_box(0); -x_646 = x_696; -x_647 = x_681; -x_648 = x_682; +x_642 = x_677; +x_643 = x_678; +x_644 = x_680; +x_645 = x_679; +x_646 = x_681; +x_647 = x_682; +x_648 = x_684; x_649 = x_683; -x_650 = x_684; +x_650 = x_686; x_651 = x_685; -x_652 = x_686; -x_653 = x_687; -x_654 = x_689; +x_652 = x_687; +x_653 = x_689; +x_654 = x_688; x_655 = x_697; goto block_666; } else { lean_object* x_698; lean_object* x_699; lean_object* x_700; -x_698 = lean_ctor_get(x_687, 0); +x_698 = lean_ctor_get(x_676, 0); x_699 = l_Lean_Elab_Command_elabInitialize___closed__41; lean_inc(x_698); x_700 = lean_array_push(x_699, x_698); -x_634 = x_668; -x_635 = x_669; -x_636 = x_670; -x_637 = x_671; +x_634 = x_667; +x_635 = x_668; +x_636 = x_696; +x_637 = lean_box(0); x_638 = x_672; x_639 = x_674; x_640 = x_675; x_641 = x_676; -x_642 = x_680; -x_643 = x_679; -x_644 = x_678; -x_645 = lean_box(0); -x_646 = x_696; -x_647 = x_681; -x_648 = x_682; +x_642 = x_677; +x_643 = x_678; +x_644 = x_680; +x_645 = x_679; +x_646 = x_681; +x_647 = x_682; +x_648 = x_684; x_649 = x_683; -x_650 = x_684; +x_650 = x_686; x_651 = x_685; -x_652 = x_686; -x_653 = x_687; -x_654 = x_689; +x_652 = x_687; +x_653 = x_689; +x_654 = x_688; x_655 = x_700; goto block_666; } @@ -20512,122 +20512,122 @@ goto block_666; block_748: { lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; -lean_inc_ref(x_712); -x_722 = l_Array_append___redArg(x_712, x_721); +lean_inc_ref(x_702); +x_722 = l_Array_append___redArg(x_702, x_721); lean_dec_ref(x_721); -lean_inc(x_715); -lean_inc(x_705); +lean_inc(x_707); +lean_inc(x_711); x_723 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_723, 0, x_705); -lean_ctor_set(x_723, 1, x_715); +lean_ctor_set(x_723, 0, x_711); +lean_ctor_set(x_723, 1, x_707); lean_ctor_set(x_723, 2, x_722); x_724 = l_Lean_Elab_Command_elabInitialize___closed__44; -lean_inc_ref(x_707); -x_725 = l_Lean_Name_mkStr4(x_5, x_6, x_707, x_724); +lean_inc_ref(x_708); +x_725 = l_Lean_Name_mkStr4(x_5, x_6, x_708, x_724); x_726 = l_Lean_Elab_Command_elabInitialize___closed__45; -lean_inc(x_705); +lean_inc(x_711); x_727 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_727, 0, x_705); +lean_ctor_set(x_727, 0, x_711); lean_ctor_set(x_727, 1, x_726); x_728 = l_Lean_Elab_Command_elabInitialize___closed__48; -lean_inc_ref(x_707); -x_729 = l_Lean_Name_mkStr4(x_5, x_6, x_707, x_728); +lean_inc_ref(x_708); +x_729 = l_Lean_Name_mkStr4(x_5, x_6, x_708, x_728); x_730 = l_Lean_Elab_Command_elabInitialize___closed__49; -lean_inc_ref(x_707); -x_731 = l_Lean_Name_mkStr4(x_5, x_6, x_707, x_730); -lean_inc_ref(x_712); -lean_inc(x_715); -lean_inc(x_705); +lean_inc_ref(x_708); +x_731 = l_Lean_Name_mkStr4(x_5, x_6, x_708, x_730); +lean_inc_ref(x_702); +lean_inc(x_707); +lean_inc(x_711); x_732 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_732, 0, x_705); -lean_ctor_set(x_732, 1, x_715); -lean_ctor_set(x_732, 2, x_712); +lean_ctor_set(x_732, 0, x_711); +lean_ctor_set(x_732, 1, x_707); +lean_ctor_set(x_732, 2, x_702); lean_inc_ref(x_732); -lean_inc(x_705); -x_733 = l_Lean_Syntax_node1(x_705, x_731, x_732); +lean_inc(x_711); +x_733 = l_Lean_Syntax_node1(x_711, x_731, x_732); x_734 = l_Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9___closed__2; x_735 = l_Lean_Elab_Command_elabInitialize___closed__8; x_736 = l_Lean_Elab_Command_elabInitialize___closed__9; -x_737 = l_Lean_addMacroScope(x_713, x_736, x_702); +x_737 = l_Lean_addMacroScope(x_715, x_736, x_703); x_738 = lean_box(0); -lean_inc(x_705); +lean_inc(x_711); x_739 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_739, 0, x_705); +lean_ctor_set(x_739, 0, x_711); lean_ctor_set(x_739, 1, x_735); lean_ctor_set(x_739, 2, x_737); lean_ctor_set(x_739, 3, x_738); -lean_inc(x_715); -lean_inc(x_705); -x_740 = l_Lean_Syntax_node1(x_705, x_715, x_739); -lean_inc(x_705); -x_741 = l_Lean_Syntax_node2(x_705, x_734, x_714, x_740); -lean_inc(x_705); -x_742 = l_Lean_Syntax_node2(x_705, x_729, x_733, x_741); +lean_inc(x_707); +lean_inc(x_711); +x_740 = l_Lean_Syntax_node1(x_711, x_707, x_739); +lean_inc(x_711); +x_741 = l_Lean_Syntax_node2(x_711, x_734, x_705, x_740); +lean_inc(x_711); +x_742 = l_Lean_Syntax_node2(x_711, x_729, x_733, x_741); x_743 = l_Lean_Elab_Command_elabInitialize___closed__46; -lean_inc(x_705); +lean_inc(x_711); x_744 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_744, 0, x_705); +lean_ctor_set(x_744, 0, x_711); lean_ctor_set(x_744, 1, x_743); x_745 = l_Array_mkArray2___redArg(x_742, x_744); -if (lean_obj_tag(x_719) == 0) +if (lean_obj_tag(x_716) == 0) { lean_object* x_746; x_746 = l_Lean_Elab_Command_elabInitialize___closed__41; -x_667 = x_745; -x_668 = x_703; -x_669 = x_732; -x_670 = x_704; -x_671 = x_705; -x_672 = x_706; -x_673 = x_725; -x_674 = x_708; +x_667 = x_702; +x_668 = x_704; +x_669 = x_727; +x_670 = lean_box(0); +x_671 = x_725; +x_672 = x_738; +x_673 = x_745; +x_674 = x_723; x_675 = x_707; x_676 = x_709; -x_677 = lean_box(0); -x_678 = x_736; -x_679 = x_738; -x_680 = x_723; -x_681 = x_711; -x_682 = x_712; -x_683 = x_735; -x_684 = x_715; -x_685 = x_716; -x_686 = x_717; -x_687 = x_718; -x_688 = x_727; -x_689 = x_720; +x_677 = x_710; +x_678 = x_708; +x_679 = x_712; +x_680 = x_711; +x_681 = x_713; +x_682 = x_714; +x_683 = x_718; +x_684 = x_717; +x_685 = x_735; +x_686 = x_719; +x_687 = x_720; +x_688 = x_736; +x_689 = x_732; x_690 = x_746; goto block_701; } else { lean_object* x_747; -x_747 = lean_ctor_get(x_719, 0); +x_747 = lean_ctor_get(x_716, 0); lean_inc(x_747); -lean_dec_ref(x_719); -x_667 = x_745; -x_668 = x_703; -x_669 = x_732; -x_670 = x_704; -x_671 = x_705; -x_672 = x_706; -x_673 = x_725; -x_674 = x_708; +lean_dec_ref(x_716); +x_667 = x_702; +x_668 = x_704; +x_669 = x_727; +x_670 = lean_box(0); +x_671 = x_725; +x_672 = x_738; +x_673 = x_745; +x_674 = x_723; x_675 = x_707; x_676 = x_709; -x_677 = lean_box(0); -x_678 = x_736; -x_679 = x_738; -x_680 = x_723; -x_681 = x_711; -x_682 = x_712; -x_683 = x_735; -x_684 = x_715; -x_685 = x_716; -x_686 = x_717; -x_687 = x_718; -x_688 = x_727; -x_689 = x_720; +x_677 = x_710; +x_678 = x_708; +x_679 = x_712; +x_680 = x_711; +x_681 = x_713; +x_682 = x_714; +x_683 = x_718; +x_684 = x_717; +x_685 = x_735; +x_686 = x_719; +x_687 = x_720; +x_688 = x_736; +x_689 = x_732; x_690 = x_747; goto block_701; } @@ -20638,24 +20638,24 @@ lean_object* x_766; lean_object* x_767; lean_object* x_768; x_766 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__4; x_767 = l_Lean_Elab_addDeclarationRangesForBuiltin___at___00Lean_Elab_Command_elabAxiom_spec__26___closed__4; x_768 = l_Lean_Elab_Command_expandNamespacedDeclaration___closed__5; -if (lean_obj_tag(x_756) == 0) +if (lean_obj_tag(x_751) == 0) { lean_object* x_769; x_769 = l_Lean_Elab_Command_elabInitialize___closed__41; -x_702 = x_749; -x_703 = x_750; -x_704 = x_766; -x_705 = x_751; -x_706 = x_752; -x_707 = x_753; -x_708 = x_754; -x_709 = x_755; -x_710 = lean_box(0); -x_711 = x_757; -x_712 = x_768; -x_713 = x_764; -x_714 = x_758; -x_715 = x_767; +x_702 = x_768; +x_703 = x_749; +x_704 = x_750; +x_705 = x_752; +x_706 = lean_box(0); +x_707 = x_767; +x_708 = x_753; +x_709 = x_754; +x_710 = x_755; +x_711 = x_756; +x_712 = x_757; +x_713 = x_758; +x_714 = x_766; +x_715 = x_764; x_716 = x_759; x_717 = x_760; x_718 = x_761; @@ -20667,24 +20667,24 @@ goto block_748; else { lean_object* x_770; lean_object* x_771; -x_770 = lean_ctor_get(x_756, 0); +x_770 = lean_ctor_get(x_751, 0); lean_inc(x_770); -lean_dec_ref(x_756); +lean_dec_ref(x_751); x_771 = l_Array_mkArray1___redArg(x_770); -x_702 = x_749; -x_703 = x_750; -x_704 = x_766; -x_705 = x_751; -x_706 = x_752; -x_707 = x_753; -x_708 = x_754; -x_709 = x_755; -x_710 = lean_box(0); -x_711 = x_757; -x_712 = x_768; -x_713 = x_764; -x_714 = x_758; -x_715 = x_767; +x_702 = x_768; +x_703 = x_749; +x_704 = x_750; +x_705 = x_752; +x_706 = lean_box(0); +x_707 = x_767; +x_708 = x_753; +x_709 = x_754; +x_710 = x_755; +x_711 = x_756; +x_712 = x_757; +x_713 = x_758; +x_714 = x_766; +x_715 = x_764; x_716 = x_759; x_717 = x_760; x_718 = x_761; @@ -20697,43 +20697,43 @@ goto block_748; block_802: { lean_object* x_787; -x_787 = l_Lean_Elab_Command_getRef___redArg(x_779); +x_787 = l_Lean_Elab_Command_getRef___redArg(x_773); if (lean_obj_tag(x_787) == 0) { lean_object* x_788; lean_object* x_789; x_788 = lean_ctor_get(x_787, 0); lean_inc(x_788); lean_dec_ref(x_787); -x_789 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_779); +x_789 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_773); if (lean_obj_tag(x_789) == 0) { lean_object* x_790; lean_object* x_791; lean_object* x_792; x_790 = lean_ctor_get(x_789, 0); lean_inc(x_790); lean_dec_ref(x_789); -x_791 = lean_ctor_get(x_779, 5); +x_791 = lean_ctor_get(x_773, 5); x_792 = l_Lean_SourceInfo_fromRef(x_788, x_777); lean_dec(x_788); if (lean_obj_tag(x_791) == 0) { lean_object* x_793; lean_object* x_794; -x_793 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_776); +x_793 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_782); x_794 = lean_ctor_get(x_793, 0); lean_inc(x_794); lean_dec_ref(x_793); x_749 = x_790; x_750 = x_773; -x_751 = x_792; -x_752 = x_774; -x_753 = x_775; -x_754 = x_776; +x_751 = x_774; +x_752 = x_775; +x_753 = x_778; +x_754 = x_786; x_755 = x_777; -x_756 = x_778; +x_756 = x_792; x_757 = x_779; x_758 = x_780; x_759 = x_781; -x_760 = x_783; -x_761 = x_786; +x_760 = x_782; +x_761 = x_783; x_762 = x_784; x_763 = x_785; x_764 = x_794; @@ -20747,17 +20747,17 @@ x_795 = lean_ctor_get(x_791, 0); lean_inc(x_795); x_749 = x_790; x_750 = x_773; -x_751 = x_792; -x_752 = x_774; -x_753 = x_775; -x_754 = x_776; +x_751 = x_774; +x_752 = x_775; +x_753 = x_778; +x_754 = x_786; x_755 = x_777; -x_756 = x_778; +x_756 = x_792; x_757 = x_779; x_758 = x_780; x_759 = x_781; -x_760 = x_783; -x_761 = x_786; +x_760 = x_782; +x_761 = x_783; x_762 = x_784; x_763 = x_785; x_764 = x_795; @@ -20773,14 +20773,14 @@ lean_dec(x_786); lean_dec(x_785); lean_dec(x_784); lean_dec(x_783); +lean_dec(x_782); lean_dec(x_781); lean_dec(x_780); -lean_dec_ref(x_779); -lean_dec(x_778); -lean_dec(x_776); -lean_dec_ref(x_775); +lean_dec(x_779); +lean_dec_ref(x_778); +lean_dec(x_775); lean_dec(x_774); -lean_dec(x_773); +lean_dec_ref(x_773); x_796 = !lean_is_exclusive(x_789); if (x_796 == 0) { @@ -20805,14 +20805,14 @@ lean_dec(x_786); lean_dec(x_785); lean_dec(x_784); lean_dec(x_783); +lean_dec(x_782); lean_dec(x_781); lean_dec(x_780); -lean_dec_ref(x_779); -lean_dec(x_778); -lean_dec(x_776); -lean_dec_ref(x_775); +lean_dec(x_779); +lean_dec_ref(x_778); +lean_dec(x_775); lean_dec(x_774); -lean_dec(x_773); +lean_dec_ref(x_773); x_799 = !lean_is_exclusive(x_787); if (x_799 == 0) { @@ -20865,17 +20865,17 @@ if (lean_obj_tag(x_822) == 0) { lean_object* x_823; x_823 = lean_box(0); -x_773 = x_804; -x_774 = x_813; -x_775 = x_810; -x_776 = x_815; +x_773 = x_814; +x_774 = x_805; +x_775 = x_806; +x_776 = lean_box(0); x_777 = x_811; -x_778 = x_812; -x_779 = x_814; -x_780 = x_805; -x_781 = x_806; -x_782 = lean_box(0); -x_783 = x_807; +x_778 = x_810; +x_779 = x_812; +x_780 = x_804; +x_781 = x_807; +x_782 = x_815; +x_783 = x_813; x_784 = x_808; x_785 = x_809; x_786 = x_823; @@ -20887,17 +20887,17 @@ uint8_t x_824; x_824 = !lean_is_exclusive(x_822); if (x_824 == 0) { -x_773 = x_804; -x_774 = x_813; -x_775 = x_810; -x_776 = x_815; +x_773 = x_814; +x_774 = x_805; +x_775 = x_806; +x_776 = lean_box(0); x_777 = x_811; -x_778 = x_812; -x_779 = x_814; -x_780 = x_805; -x_781 = x_806; -x_782 = lean_box(0); -x_783 = x_807; +x_778 = x_810; +x_779 = x_812; +x_780 = x_804; +x_781 = x_807; +x_782 = x_815; +x_783 = x_813; x_784 = x_808; x_785 = x_809; x_786 = x_822; @@ -20911,17 +20911,17 @@ lean_inc(x_825); lean_dec(x_822); x_826 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_826, 0, x_825); -x_773 = x_804; -x_774 = x_813; -x_775 = x_810; -x_776 = x_815; +x_773 = x_814; +x_774 = x_805; +x_775 = x_806; +x_776 = lean_box(0); x_777 = x_811; -x_778 = x_812; -x_779 = x_814; -x_780 = x_805; -x_781 = x_806; -x_782 = lean_box(0); -x_783 = x_807; +x_778 = x_810; +x_779 = x_812; +x_780 = x_804; +x_781 = x_807; +x_782 = x_815; +x_783 = x_813; x_784 = x_808; x_785 = x_809; x_786 = x_826; @@ -20947,8 +20947,8 @@ lean_object* x_845; lean_object* x_846; lean_dec(x_842); lean_dec(x_837); lean_dec(x_836); +lean_dec(x_835); lean_dec_ref(x_834); -lean_dec(x_833); lean_dec(x_832); lean_dec(x_831); lean_dec(x_830); @@ -20974,8 +20974,8 @@ lean_object* x_850; lean_object* x_851; lean_dec(x_847); lean_dec(x_837); lean_dec(x_836); +lean_dec(x_835); lean_dec_ref(x_834); -lean_dec(x_833); lean_dec(x_832); lean_dec(x_831); lean_dec(x_830); @@ -20999,10 +20999,10 @@ x_804 = x_829; x_805 = x_830; x_806 = x_831; x_807 = x_832; -x_808 = x_833; -x_809 = x_837; +x_808 = x_837; +x_809 = x_835; x_810 = x_834; -x_811 = x_835; +x_811 = x_833; x_812 = x_836; x_813 = x_853; x_814 = x_838; @@ -21022,10 +21022,10 @@ x_804 = x_829; x_805 = x_830; x_806 = x_831; x_807 = x_832; -x_808 = x_833; -x_809 = x_837; +x_808 = x_837; +x_809 = x_835; x_810 = x_834; -x_811 = x_835; +x_811 = x_833; x_812 = x_836; x_813 = x_854; x_814 = x_838; @@ -21037,15 +21037,15 @@ goto block_827; block_887: { lean_object* x_868; uint8_t x_869; -x_868 = l_Lean_Syntax_getArg(x_12, x_859); +x_868 = l_Lean_Syntax_getArg(x_12, x_862); x_869 = l_Lean_Syntax_matchesNull(x_868, x_11); if (x_869 == 0) { lean_object* x_870; lean_object* x_871; lean_dec(x_864); lean_dec(x_863); -lean_dec_ref(x_861); -lean_dec(x_860); +lean_dec_ref(x_860); +lean_dec(x_859); lean_dec(x_858); lean_dec(x_857); lean_dec(x_856); @@ -21074,8 +21074,8 @@ lean_dec(x_874); lean_dec(x_872); lean_dec(x_864); lean_dec(x_863); -lean_dec_ref(x_861); -lean_dec(x_860); +lean_dec_ref(x_860); +lean_dec(x_859); lean_dec(x_858); lean_dec(x_857); lean_dec(x_856); @@ -21100,8 +21100,8 @@ lean_dec(x_879); lean_dec(x_872); lean_dec(x_864); lean_dec(x_863); -lean_dec_ref(x_861); -lean_dec(x_860); +lean_dec_ref(x_860); +lean_dec(x_859); lean_dec(x_858); lean_dec(x_857); lean_dec(x_856); @@ -21122,10 +21122,10 @@ x_828 = x_872; x_829 = x_856; x_830 = x_857; x_831 = x_858; -x_832 = x_860; -x_833 = x_864; -x_834 = x_861; -x_835 = x_862; +x_832 = x_864; +x_833 = x_861; +x_834 = x_860; +x_835 = x_859; x_836 = x_863; x_837 = x_885; x_838 = x_865; @@ -21144,10 +21144,10 @@ x_828 = x_872; x_829 = x_856; x_830 = x_857; x_831 = x_858; -x_832 = x_860; -x_833 = x_864; -x_834 = x_861; -x_835 = x_862; +x_832 = x_864; +x_833 = x_861; +x_834 = x_860; +x_835 = x_859; x_836 = x_863; x_837 = x_886; x_838 = x_865; @@ -21172,9 +21172,9 @@ if (x_901 == 0) lean_object* x_902; lean_object* x_903; lean_dec(x_899); lean_dec(x_895); -lean_dec_ref(x_893); -lean_dec(x_892); -lean_dec(x_890); +lean_dec(x_894); +lean_dec(x_893); +lean_dec_ref(x_892); lean_dec(x_889); lean_dec(x_888); x_902 = l_Lean_Elab_Command_elabInitialize___closed__53; @@ -21189,8 +21189,8 @@ lean_object* x_904; lean_object* x_905; lean_object* x_906; uint8_t x_907; x_904 = l_Lean_Syntax_getArg(x_899, x_11); lean_dec(x_899); x_905 = l_Lean_Elab_Command_elabInitialize___closed__44; -lean_inc_ref(x_893); -x_906 = l_Lean_Name_mkStr4(x_5, x_6, x_893, x_905); +lean_inc_ref(x_892); +x_906 = l_Lean_Name_mkStr4(x_5, x_6, x_892, x_905); lean_inc(x_904); x_907 = l_Lean_Syntax_isOfKind(x_904, x_906); lean_dec(x_906); @@ -21199,9 +21199,9 @@ if (x_907 == 0) lean_object* x_908; lean_object* x_909; lean_dec(x_904); lean_dec(x_895); -lean_dec_ref(x_893); -lean_dec(x_892); -lean_dec(x_890); +lean_dec(x_894); +lean_dec(x_893); +lean_dec_ref(x_892); lean_dec(x_889); lean_dec(x_888); x_908 = l_Lean_Elab_Command_elabInitialize___closed__53; @@ -21220,13 +21220,13 @@ lean_dec(x_910); x_912 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_912, 0, x_911); x_856 = x_888; -x_857 = x_889; -x_858 = x_890; -x_859 = x_891; +x_857 = x_895; +x_858 = x_889; +x_859 = x_893; x_860 = x_892; -x_861 = x_893; -x_862 = x_894; -x_863 = x_895; +x_861 = x_891; +x_862 = x_890; +x_863 = x_894; x_864 = x_912; x_865 = x_896; x_866 = x_897; @@ -21241,13 +21241,13 @@ lean_object* x_913; lean_dec(x_899); x_913 = lean_box(0); x_856 = x_888; -x_857 = x_889; -x_858 = x_890; -x_859 = x_891; +x_857 = x_895; +x_858 = x_889; +x_859 = x_893; x_860 = x_892; -x_861 = x_893; -x_862 = x_894; -x_863 = x_895; +x_861 = x_891; +x_862 = x_890; +x_863 = x_894; x_864 = x_913; x_865 = x_896; x_866 = x_897; @@ -21263,29 +21263,29 @@ x_925 = l_Lean_mkIdentFrom(x_1, x_923, x_924); lean_dec(x_1); if (lean_obj_tag(x_918) == 0) { -lean_dec(x_916); +lean_dec(x_917); x_416 = x_925; -x_417 = x_919; -x_418 = x_920; -x_419 = x_922; -x_420 = x_924; -x_421 = x_921; -x_422 = x_917; +x_417 = x_920; +x_418 = x_924; +x_419 = x_921; +x_420 = x_922; +x_421 = x_916; +x_422 = x_919; x_423 = lean_box(0); goto block_438; } else { -if (lean_obj_tag(x_916) == 0) +if (lean_obj_tag(x_917) == 0) { lean_dec_ref(x_918); x_416 = x_925; -x_417 = x_919; -x_418 = x_920; -x_419 = x_922; -x_420 = x_924; -x_421 = x_921; -x_422 = x_917; +x_417 = x_920; +x_418 = x_924; +x_419 = x_921; +x_420 = x_922; +x_421 = x_916; +x_422 = x_919; x_423 = lean_box(0); goto block_438; } @@ -21295,13 +21295,13 @@ if (x_151 == 0) { lean_object* x_926; lean_object* x_927; lean_dec(x_925); -lean_dec_ref(x_922); -lean_dec(x_920); +lean_dec(x_922); +lean_dec_ref(x_921); lean_dec_ref(x_918); -lean_dec_ref(x_916); +lean_dec_ref(x_917); x_926 = l_Lean_Elab_Command_elabInitialize___closed__53; -x_927 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_926, x_921, x_917); -lean_dec(x_917); +x_927 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_926, x_916, x_919); +lean_dec(x_919); lean_dec(x_12); return x_927; } @@ -21311,11 +21311,11 @@ lean_object* x_928; uint8_t x_929; x_928 = lean_ctor_get(x_918, 0); lean_inc(x_928); lean_dec_ref(x_918); -x_929 = !lean_is_exclusive(x_916); +x_929 = !lean_is_exclusive(x_917); if (x_929 == 0) { lean_object* x_930; lean_object* x_931; uint8_t x_932; -x_930 = lean_ctor_get(x_916, 0); +x_930 = lean_ctor_get(x_917, 0); x_931 = l_Lean_Syntax_getArg(x_12, x_11); x_932 = l_Lean_Syntax_isNone(x_931); if (x_932 == 0) @@ -21327,15 +21327,15 @@ if (x_933 == 0) { lean_object* x_934; lean_object* x_935; lean_dec(x_931); -lean_free_object(x_916); +lean_free_object(x_917); lean_dec(x_930); lean_dec(x_928); lean_dec(x_925); -lean_dec_ref(x_922); -lean_dec(x_920); +lean_dec(x_922); +lean_dec_ref(x_921); x_934 = l_Lean_Elab_Command_elabInitialize___closed__53; -x_935 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_934, x_921, x_917); -lean_dec(x_917); +x_935 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_934, x_916, x_919); +lean_dec(x_919); lean_dec(x_12); return x_935; } @@ -21351,31 +21351,31 @@ if (x_938 == 0) { lean_object* x_939; lean_object* x_940; lean_dec(x_936); -lean_free_object(x_916); +lean_free_object(x_917); lean_dec(x_930); lean_dec(x_928); lean_dec(x_925); -lean_dec_ref(x_922); -lean_dec(x_920); +lean_dec(x_922); +lean_dec_ref(x_921); x_939 = l_Lean_Elab_Command_elabInitialize___closed__53; -x_940 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_939, x_921, x_917); -lean_dec(x_917); +x_940 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_939, x_916, x_919); +lean_dec(x_919); lean_dec(x_12); return x_940; } else { -lean_ctor_set(x_916, 0, x_936); -x_888 = x_928; +lean_ctor_set(x_917, 0, x_936); +x_888 = x_930; x_889 = x_925; -x_890 = x_930; -x_891 = x_919; -x_892 = x_920; -x_893 = x_922; -x_894 = x_924; -x_895 = x_916; -x_896 = x_921; -x_897 = x_917; +x_890 = x_920; +x_891 = x_924; +x_892 = x_921; +x_893 = x_928; +x_894 = x_922; +x_895 = x_917; +x_896 = x_916; +x_897 = x_919; x_898 = lean_box(0); goto block_914; } @@ -21385,18 +21385,18 @@ else { lean_object* x_941; lean_dec(x_931); -lean_free_object(x_916); +lean_free_object(x_917); x_941 = lean_box(0); -x_888 = x_928; +x_888 = x_930; x_889 = x_925; -x_890 = x_930; -x_891 = x_919; -x_892 = x_920; -x_893 = x_922; -x_894 = x_924; +x_890 = x_920; +x_891 = x_924; +x_892 = x_921; +x_893 = x_928; +x_894 = x_922; x_895 = x_941; -x_896 = x_921; -x_897 = x_917; +x_896 = x_916; +x_897 = x_919; x_898 = lean_box(0); goto block_914; } @@ -21404,9 +21404,9 @@ goto block_914; else { lean_object* x_942; lean_object* x_943; uint8_t x_944; -x_942 = lean_ctor_get(x_916, 0); +x_942 = lean_ctor_get(x_917, 0); lean_inc(x_942); -lean_dec(x_916); +lean_dec(x_917); x_943 = l_Lean_Syntax_getArg(x_12, x_11); x_944 = l_Lean_Syntax_isNone(x_943); if (x_944 == 0) @@ -21421,11 +21421,11 @@ lean_dec(x_943); lean_dec(x_942); lean_dec(x_928); lean_dec(x_925); -lean_dec_ref(x_922); -lean_dec(x_920); +lean_dec(x_922); +lean_dec_ref(x_921); x_946 = l_Lean_Elab_Command_elabInitialize___closed__53; -x_947 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_946, x_921, x_917); -lean_dec(x_917); +x_947 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_946, x_916, x_919); +lean_dec(x_919); lean_dec(x_12); return x_947; } @@ -21444,11 +21444,11 @@ lean_dec(x_948); lean_dec(x_942); lean_dec(x_928); lean_dec(x_925); -lean_dec_ref(x_922); -lean_dec(x_920); +lean_dec(x_922); +lean_dec_ref(x_921); x_951 = l_Lean_Elab_Command_elabInitialize___closed__53; -x_952 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_951, x_921, x_917); -lean_dec(x_917); +x_952 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_951, x_916, x_919); +lean_dec(x_919); lean_dec(x_12); return x_952; } @@ -21457,16 +21457,16 @@ else lean_object* x_953; x_953 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_953, 0, x_948); -x_888 = x_928; +x_888 = x_942; x_889 = x_925; -x_890 = x_942; -x_891 = x_919; -x_892 = x_920; -x_893 = x_922; -x_894 = x_924; +x_890 = x_920; +x_891 = x_924; +x_892 = x_921; +x_893 = x_928; +x_894 = x_922; x_895 = x_953; -x_896 = x_921; -x_897 = x_917; +x_896 = x_916; +x_897 = x_919; x_898 = lean_box(0); goto block_914; } @@ -21477,16 +21477,16 @@ else lean_object* x_954; lean_dec(x_943); x_954 = lean_box(0); -x_888 = x_928; +x_888 = x_942; x_889 = x_925; -x_890 = x_942; -x_891 = x_919; -x_892 = x_920; -x_893 = x_922; -x_894 = x_924; +x_890 = x_920; +x_891 = x_924; +x_892 = x_921; +x_893 = x_928; +x_894 = x_922; x_895 = x_954; -x_896 = x_921; -x_897 = x_917; +x_896 = x_916; +x_897 = x_919; x_898 = lean_box(0); goto block_914; } @@ -21509,13 +21509,13 @@ if (x_965 == 0) lean_object* x_966; x_966 = l_Lean_Elab_Command_elabInitialize___closed__59; x_915 = lean_box(0); -x_916 = x_957; -x_917 = x_959; +x_916 = x_958; +x_917 = x_957; x_918 = x_956; -x_919 = x_961; -x_920 = x_962; -x_921 = x_958; -x_922 = x_963; +x_919 = x_959; +x_920 = x_961; +x_921 = x_963; +x_922 = x_962; x_923 = x_966; goto block_955; } @@ -21524,13 +21524,13 @@ else lean_object* x_967; x_967 = l_Lean_Elab_Command_elabInitialize___closed__51; x_915 = lean_box(0); -x_916 = x_957; -x_917 = x_959; +x_916 = x_958; +x_917 = x_957; x_918 = x_956; -x_919 = x_961; -x_920 = x_962; -x_921 = x_958; -x_922 = x_963; +x_919 = x_959; +x_920 = x_961; +x_921 = x_963; +x_922 = x_962; x_923 = x_967; goto block_955; } @@ -21539,14 +21539,14 @@ goto block_955; block_184: { lean_object* x_174; lean_object* x_175; -lean_inc_ref(x_159); -x_174 = l_Array_append___redArg(x_159, x_173); +lean_inc_ref(x_161); +x_174 = l_Array_append___redArg(x_161, x_173); lean_dec_ref(x_173); -lean_inc(x_164); -lean_inc(x_167); +lean_inc(x_166); +lean_inc(x_157); x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_167); -lean_ctor_set(x_175, 1, x_164); +lean_ctor_set(x_175, 0, x_157); +lean_ctor_set(x_175, 1, x_166); lean_ctor_set(x_175, 2, x_174); if (lean_obj_tag(x_172) == 0) { @@ -21558,15 +21558,15 @@ x_16 = x_159; x_17 = x_160; x_18 = x_161; x_19 = x_162; -x_20 = x_175; -x_21 = x_163; -x_22 = lean_box(0); -x_23 = x_164; -x_24 = x_166; -x_25 = x_167; -x_26 = x_168; -x_27 = x_169; -x_28 = x_170; +x_20 = lean_box(0); +x_21 = x_164; +x_22 = x_165; +x_23 = x_166; +x_24 = x_167; +x_25 = x_168; +x_26 = x_169; +x_27 = x_170; +x_28 = x_175; x_29 = x_171; x_30 = x_176; goto block_79; @@ -21584,8 +21584,8 @@ lean_dec(x_177); x_181 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_181, 0, x_180); lean_ctor_set(x_181, 1, x_178); -lean_inc(x_167); -x_182 = l_Lean_Syntax_node1(x_167, x_179, x_181); +lean_inc(x_157); +x_182 = l_Lean_Syntax_node1(x_157, x_179, x_181); x_183 = l_Array_mkArray1___redArg(x_182); x_14 = x_157; x_15 = x_158; @@ -21593,15 +21593,15 @@ x_16 = x_159; x_17 = x_160; x_18 = x_161; x_19 = x_162; -x_20 = x_175; -x_21 = x_163; -x_22 = lean_box(0); -x_23 = x_164; -x_24 = x_166; -x_25 = x_167; -x_26 = x_168; -x_27 = x_169; -x_28 = x_170; +x_20 = lean_box(0); +x_21 = x_164; +x_22 = x_165; +x_23 = x_166; +x_24 = x_167; +x_25 = x_168; +x_26 = x_169; +x_27 = x_170; +x_28 = x_175; x_29 = x_171; x_30 = x_183; goto block_79; @@ -21613,84 +21613,84 @@ lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_inc_ref(x_187); x_201 = l_Array_append___redArg(x_187, x_200); lean_dec_ref(x_200); -lean_inc(x_190); -lean_inc(x_193); +lean_inc(x_191); +lean_inc(x_185); x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_193); -lean_ctor_set(x_202, 1, x_190); +lean_ctor_set(x_202, 0, x_185); +lean_ctor_set(x_202, 1, x_191); lean_ctor_set(x_202, 2, x_201); x_203 = l_Lean_Elab_Command_elabInitialize___closed__44; -lean_inc_ref(x_194); -x_204 = l_Lean_Name_mkStr4(x_5, x_6, x_194, x_203); +lean_inc_ref(x_192); +x_204 = l_Lean_Name_mkStr4(x_5, x_6, x_192, x_203); x_205 = l_Lean_Elab_Command_elabInitialize___closed__45; -lean_inc(x_193); +lean_inc(x_185); x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_193); +lean_ctor_set(x_206, 0, x_185); lean_ctor_set(x_206, 1, x_205); -x_207 = lean_array_size(x_196); +x_207 = lean_array_size(x_197); x_208 = 0; -x_209 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabInitialize_spec__0(x_207, x_208, x_196); +x_209 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabInitialize_spec__0(x_207, x_208, x_197); x_210 = l_Lean_Elab_Command_elabInitialize___closed__47; x_211 = l_Lean_mkSepArray(x_209, x_210); lean_dec_ref(x_209); lean_inc_ref(x_187); x_212 = l_Array_append___redArg(x_187, x_211); lean_dec_ref(x_211); -lean_inc(x_190); -lean_inc(x_193); +lean_inc(x_191); +lean_inc(x_185); x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_193); -lean_ctor_set(x_213, 1, x_190); +lean_ctor_set(x_213, 0, x_185); +lean_ctor_set(x_213, 1, x_191); lean_ctor_set(x_213, 2, x_212); x_214 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Lean_Elab_Command_elabAttr_spec__0___closed__4; -lean_inc(x_193); +lean_inc(x_185); x_215 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_215, 0, x_193); +lean_ctor_set(x_215, 0, x_185); lean_ctor_set(x_215, 1, x_214); -lean_inc(x_193); -x_216 = l_Lean_Syntax_node3(x_193, x_204, x_206, x_213, x_215); -lean_inc(x_190); -lean_inc(x_193); -x_217 = l_Lean_Syntax_node1(x_193, x_190, x_216); -lean_inc(x_190); -lean_inc(x_193); -x_218 = l_Lean_Syntax_node1(x_193, x_190, x_199); +lean_inc(x_185); +x_216 = l_Lean_Syntax_node3(x_185, x_204, x_206, x_213, x_215); +lean_inc(x_191); +lean_inc(x_185); +x_217 = l_Lean_Syntax_node1(x_185, x_191, x_216); +lean_inc(x_191); +lean_inc(x_185); +x_218 = l_Lean_Syntax_node1(x_185, x_191, x_195); lean_inc_ref(x_187); -lean_inc(x_190); -lean_inc(x_193); +lean_inc(x_191); +lean_inc(x_185); x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_193); -lean_ctor_set(x_219, 1, x_190); +lean_ctor_set(x_219, 0, x_185); +lean_ctor_set(x_219, 1, x_191); lean_ctor_set(x_219, 2, x_187); -if (lean_obj_tag(x_192) == 0) +if (lean_obj_tag(x_194) == 0) { lean_object* x_220; x_220 = l_Lean_Elab_Command_elabInitialize___closed__41; x_157 = x_185; x_158 = x_186; -x_159 = x_187; +x_159 = x_218; x_160 = x_217; -x_161 = x_188; -x_162 = x_189; -x_163 = x_202; -x_164 = x_190; -x_165 = lean_box(0); -x_166 = x_218; -x_167 = x_193; -x_168 = x_195; -x_169 = x_194; -x_170 = x_219; -x_171 = x_197; -x_172 = x_198; +x_161 = x_187; +x_162 = x_219; +x_163 = lean_box(0); +x_164 = x_189; +x_165 = x_190; +x_166 = x_191; +x_167 = x_192; +x_168 = x_193; +x_169 = x_196; +x_170 = x_202; +x_171 = x_198; +x_172 = x_199; x_173 = x_220; goto block_184; } else { lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_221 = lean_ctor_get(x_192, 0); +x_221 = lean_ctor_get(x_194, 0); lean_inc(x_221); -lean_dec_ref(x_192); +lean_dec_ref(x_194); x_222 = l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__12_spec__12___closed__16; x_223 = l_Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4___closed__7; x_224 = l_Lean_SourceInfo_fromRef(x_221, x_156); @@ -21698,25 +21698,25 @@ lean_dec(x_221); x_225 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_225, 0, x_224); lean_ctor_set(x_225, 1, x_222); -lean_inc(x_193); -x_226 = l_Lean_Syntax_node1(x_193, x_223, x_225); +lean_inc(x_185); +x_226 = l_Lean_Syntax_node1(x_185, x_223, x_225); x_227 = l_Array_mkArray1___redArg(x_226); x_157 = x_185; x_158 = x_186; -x_159 = x_187; +x_159 = x_218; x_160 = x_217; -x_161 = x_188; -x_162 = x_189; -x_163 = x_202; -x_164 = x_190; -x_165 = lean_box(0); -x_166 = x_218; -x_167 = x_193; -x_168 = x_195; -x_169 = x_194; -x_170 = x_219; -x_171 = x_197; -x_172 = x_198; +x_161 = x_187; +x_162 = x_219; +x_163 = lean_box(0); +x_164 = x_189; +x_165 = x_190; +x_166 = x_191; +x_167 = x_192; +x_168 = x_193; +x_169 = x_196; +x_170 = x_202; +x_171 = x_198; +x_172 = x_199; x_173 = x_227; goto block_184; } @@ -21725,24 +21725,24 @@ goto block_184; { lean_object* x_244; x_244 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__4; -if (lean_obj_tag(x_237) == 0) +if (lean_obj_tag(x_235) == 0) { lean_object* x_245; x_245 = l_Lean_Elab_Command_elabInitialize___closed__41; -x_185 = x_242; -x_186 = x_229; -x_187 = x_230; -x_188 = x_231; -x_189 = x_244; -x_190 = x_232; -x_191 = lean_box(0); -x_192 = x_233; -x_193 = x_234; -x_194 = x_235; -x_195 = x_236; -x_196 = x_238; -x_197 = x_239; -x_198 = x_240; +x_185 = x_229; +x_186 = x_230; +x_187 = x_231; +x_188 = lean_box(0); +x_189 = x_232; +x_190 = x_233; +x_191 = x_234; +x_192 = x_236; +x_193 = x_237; +x_194 = x_238; +x_195 = x_239; +x_196 = x_244; +x_197 = x_240; +x_198 = x_242; x_199 = x_241; x_200 = x_245; goto block_228; @@ -21750,24 +21750,24 @@ goto block_228; else { lean_object* x_246; lean_object* x_247; -x_246 = lean_ctor_get(x_237, 0); +x_246 = lean_ctor_get(x_235, 0); lean_inc(x_246); -lean_dec_ref(x_237); +lean_dec_ref(x_235); x_247 = l_Array_mkArray1___redArg(x_246); -x_185 = x_242; -x_186 = x_229; -x_187 = x_230; -x_188 = x_231; -x_189 = x_244; -x_190 = x_232; -x_191 = lean_box(0); -x_192 = x_233; -x_193 = x_234; -x_194 = x_235; -x_195 = x_236; -x_196 = x_238; -x_197 = x_239; -x_198 = x_240; +x_185 = x_229; +x_186 = x_230; +x_187 = x_231; +x_188 = lean_box(0); +x_189 = x_232; +x_190 = x_233; +x_191 = x_234; +x_192 = x_236; +x_193 = x_237; +x_194 = x_238; +x_195 = x_239; +x_196 = x_244; +x_197 = x_240; +x_198 = x_242; x_199 = x_241; x_200 = x_247; goto block_228; @@ -21776,21 +21776,21 @@ goto block_228; block_292: { lean_object* x_261; -x_261 = l_Lean_Elab_Command_getRef___redArg(x_250); +x_261 = l_Lean_Elab_Command_getRef___redArg(x_254); if (lean_obj_tag(x_261) == 0) { lean_object* x_262; lean_object* x_263; x_262 = lean_ctor_get(x_261, 0); lean_inc(x_262); lean_dec_ref(x_261); -x_263 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_250); +x_263 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_254); if (lean_obj_tag(x_263) == 0) { lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; 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; uint8_t x_280; lean_object* x_281; lean_object* x_282; x_264 = lean_ctor_get(x_263, 0); lean_inc(x_264); lean_dec_ref(x_263); -x_265 = lean_ctor_get(x_250, 5); +x_265 = lean_ctor_get(x_254, 5); x_266 = l_Lean_Elab_Command_elabInitialize___closed__48; lean_inc_ref(x_258); x_267 = l_Lean_Name_mkStr4(x_5, x_6, x_258, x_266); @@ -21799,48 +21799,48 @@ lean_inc_ref(x_258); x_269 = l_Lean_Name_mkStr4(x_5, x_6, x_258, x_268); x_270 = l_Lean_Elab_addDeclarationRangesForBuiltin___at___00Lean_Elab_Command_elabAxiom_spec__26___closed__4; x_271 = l_Lean_Elab_Command_expandNamespacedDeclaration___closed__5; -lean_inc(x_256); +lean_inc(x_251); x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_256); +lean_ctor_set(x_272, 0, x_251); lean_ctor_set(x_272, 1, x_270); lean_ctor_set(x_272, 2, x_271); lean_inc_ref(x_272); -lean_inc(x_256); -x_273 = l_Lean_Syntax_node1(x_256, x_269, x_272); +lean_inc(x_251); +x_273 = l_Lean_Syntax_node1(x_251, x_269, x_272); x_274 = l_Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9___closed__2; -lean_inc(x_252); -lean_inc(x_256); -x_275 = l_Lean_Syntax_node2(x_256, x_274, x_252, x_272); -x_276 = l_Lean_Syntax_node2(x_256, x_267, x_273, x_275); -x_277 = lean_array_push(x_257, x_276); -x_278 = l_Lean_TSyntax_getId(x_252); -lean_dec(x_252); +lean_inc(x_253); +lean_inc(x_251); +x_275 = l_Lean_Syntax_node2(x_251, x_274, x_253, x_272); +x_276 = l_Lean_Syntax_node2(x_251, x_267, x_273, x_275); +x_277 = lean_array_push(x_252, x_276); +x_278 = l_Lean_TSyntax_getId(x_253); +lean_dec(x_253); x_279 = l_Lean_Elab_Command_elabInitialize___closed__51; x_280 = lean_name_eq(x_278, x_279); lean_dec(x_278); x_281 = l_Lean_Parser_Command_visibility_ofBool(x_280); -x_282 = l_Lean_SourceInfo_fromRef(x_262, x_259); +x_282 = l_Lean_SourceInfo_fromRef(x_262, x_257); lean_dec(x_262); if (lean_obj_tag(x_265) == 0) { lean_object* x_283; lean_object* x_284; -x_283 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_251); +x_283 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_250); x_284 = lean_ctor_get(x_283, 0); lean_inc(x_284); lean_dec_ref(x_283); -x_229 = x_250; -x_230 = x_271; -x_231 = x_251; -x_232 = x_270; -x_233 = x_255; -x_234 = x_282; -x_235 = x_258; -x_236 = x_264; -x_237 = x_249; -x_238 = x_277; -x_239 = x_253; -x_240 = x_254; -x_241 = x_281; +x_229 = x_282; +x_230 = x_250; +x_231 = x_271; +x_232 = x_264; +x_233 = x_254; +x_234 = x_270; +x_235 = x_255; +x_236 = x_258; +x_237 = x_259; +x_238 = x_249; +x_239 = x_281; +x_240 = x_277; +x_241 = x_256; x_242 = x_284; x_243 = lean_box(0); goto block_248; @@ -21850,19 +21850,19 @@ else lean_object* x_285; x_285 = lean_ctor_get(x_265, 0); lean_inc(x_285); -x_229 = x_250; -x_230 = x_271; -x_231 = x_251; -x_232 = x_270; -x_233 = x_255; -x_234 = x_282; -x_235 = x_258; -x_236 = x_264; -x_237 = x_249; -x_238 = x_277; -x_239 = x_253; -x_240 = x_254; -x_241 = x_281; +x_229 = x_282; +x_230 = x_250; +x_231 = x_271; +x_232 = x_264; +x_233 = x_254; +x_234 = x_270; +x_235 = x_255; +x_236 = x_258; +x_237 = x_259; +x_238 = x_249; +x_239 = x_281; +x_240 = x_277; +x_241 = x_256; x_242 = x_285; x_243 = lean_box(0); goto block_248; @@ -21872,15 +21872,15 @@ else { uint8_t x_286; lean_dec(x_262); +lean_dec(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); lean_dec(x_256); lean_dec(x_255); -lean_dec(x_254); +lean_dec_ref(x_254); lean_dec(x_253); -lean_dec(x_252); +lean_dec_ref(x_252); lean_dec(x_251); -lean_dec_ref(x_250); +lean_dec(x_250); lean_dec(x_249); x_286 = !lean_is_exclusive(x_263); if (x_286 == 0) @@ -21902,15 +21902,15 @@ return x_288; else { uint8_t x_289; +lean_dec(x_259); lean_dec_ref(x_258); -lean_dec_ref(x_257); lean_dec(x_256); lean_dec(x_255); -lean_dec(x_254); +lean_dec_ref(x_254); lean_dec(x_253); -lean_dec(x_252); +lean_dec_ref(x_252); lean_dec(x_251); -lean_dec_ref(x_250); +lean_dec(x_250); lean_dec(x_249); x_289 = !lean_is_exclusive(x_261); if (x_289 == 0) @@ -21932,36 +21932,36 @@ return x_291; block_316: { lean_object* x_304; -x_304 = l_Lean_Elab_Command_getRef___redArg(x_294); +x_304 = l_Lean_Elab_Command_getRef___redArg(x_297); if (lean_obj_tag(x_304) == 0) { lean_object* x_305; lean_object* x_306; x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); lean_dec_ref(x_304); -x_306 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_294); +x_306 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_297); if (lean_obj_tag(x_306) == 0) { lean_object* x_307; lean_object* x_308; lean_dec_ref(x_306); -x_307 = lean_ctor_get(x_294, 5); -x_308 = l_Lean_SourceInfo_fromRef(x_305, x_302); +x_307 = lean_ctor_get(x_297, 5); +x_308 = l_Lean_SourceInfo_fromRef(x_305, x_301); lean_dec(x_305); if (lean_obj_tag(x_307) == 0) { lean_object* x_309; -x_309 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_296); +x_309 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_294); lean_dec_ref(x_309); x_249 = x_293; x_250 = x_294; -x_251 = x_296; -x_252 = x_295; -x_253 = x_298; -x_254 = x_300; -x_255 = x_299; -x_256 = x_308; -x_257 = x_303; -x_258 = x_301; +x_251 = x_308; +x_252 = x_303; +x_253 = x_296; +x_254 = x_297; +x_255 = x_298; +x_256 = x_300; +x_257 = x_301; +x_258 = x_299; x_259 = x_302; x_260 = lean_box(0); goto block_292; @@ -21970,14 +21970,14 @@ else { x_249 = x_293; x_250 = x_294; -x_251 = x_296; -x_252 = x_295; -x_253 = x_298; -x_254 = x_300; -x_255 = x_299; -x_256 = x_308; -x_257 = x_303; -x_258 = x_301; +x_251 = x_308; +x_252 = x_303; +x_253 = x_296; +x_254 = x_297; +x_255 = x_298; +x_256 = x_300; +x_257 = x_301; +x_258 = x_299; x_259 = x_302; x_260 = lean_box(0); goto block_292; @@ -21988,13 +21988,13 @@ else uint8_t x_310; lean_dec(x_305); lean_dec_ref(x_303); -lean_dec_ref(x_301); +lean_dec(x_302); lean_dec(x_300); -lean_dec(x_299); +lean_dec_ref(x_299); lean_dec(x_298); +lean_dec_ref(x_297); lean_dec(x_296); -lean_dec(x_295); -lean_dec_ref(x_294); +lean_dec(x_294); lean_dec(x_293); x_310 = !lean_is_exclusive(x_306); if (x_310 == 0) @@ -22017,13 +22017,13 @@ else { uint8_t x_313; lean_dec_ref(x_303); -lean_dec_ref(x_301); +lean_dec(x_302); lean_dec(x_300); -lean_dec(x_299); +lean_dec_ref(x_299); lean_dec(x_298); +lean_dec_ref(x_297); lean_dec(x_296); -lean_dec(x_295); -lean_dec_ref(x_294); +lean_dec(x_294); lean_dec(x_293); x_313 = !lean_is_exclusive(x_304); if (x_313 == 0) @@ -22052,8 +22052,8 @@ if (x_330 == 0) { lean_object* x_331; lean_object* x_332; lean_dec(x_324); +lean_dec(x_323); lean_dec_ref(x_322); -lean_dec(x_321); lean_dec(x_320); lean_dec(x_319); lean_dec(x_318); @@ -22067,19 +22067,19 @@ return x_332; else { lean_dec(x_12); -if (lean_obj_tag(x_318) == 0) +if (lean_obj_tag(x_319) == 0) { lean_object* x_333; x_333 = l_Lean_Elab_Command_expandMutualElement___closed__0; x_293 = x_317; -x_294 = x_325; -x_295 = x_319; -x_296 = x_326; -x_297 = lean_box(0); +x_294 = x_326; +x_295 = lean_box(0); +x_296 = x_318; +x_297 = x_325; x_298 = x_320; -x_299 = x_321; +x_299 = x_322; x_300 = x_324; -x_301 = x_322; +x_301 = x_321; x_302 = x_323; x_303 = x_333; goto block_316; @@ -22087,20 +22087,20 @@ goto block_316; else { lean_object* x_334; lean_object* x_335; -x_334 = lean_ctor_get(x_318, 0); +x_334 = lean_ctor_get(x_319, 0); lean_inc(x_334); -lean_dec_ref(x_318); +lean_dec_ref(x_319); x_335 = l_Lean_Syntax_TSepArray_getElems___redArg(x_334); lean_dec(x_334); x_293 = x_317; -x_294 = x_325; -x_295 = x_319; -x_296 = x_326; -x_297 = lean_box(0); +x_294 = x_326; +x_295 = lean_box(0); +x_296 = x_318; +x_297 = x_325; x_298 = x_320; -x_299 = x_321; +x_299 = x_322; x_300 = x_324; -x_301 = x_322; +x_301 = x_321; x_302 = x_323; x_303 = x_335; goto block_316; @@ -22123,8 +22123,8 @@ if (x_350 == 0) lean_object* x_351; lean_object* x_352; lean_dec(x_348); lean_dec(x_343); -lean_dec_ref(x_341); -lean_dec(x_340); +lean_dec(x_342); +lean_dec_ref(x_340); lean_dec(x_339); lean_dec(x_338); lean_dec(x_337); @@ -22147,8 +22147,8 @@ if (x_355 == 0) lean_object* x_356; lean_object* x_357; lean_dec(x_353); lean_dec(x_343); -lean_dec_ref(x_341); -lean_dec(x_340); +lean_dec(x_342); +lean_dec_ref(x_340); lean_dec(x_339); lean_dec(x_338); lean_dec(x_337); @@ -22165,12 +22165,12 @@ x_358 = l_Lean_Syntax_getArg(x_353, x_11); lean_dec(x_353); x_359 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_359, 0, x_358); -x_317 = x_337; -x_318 = x_338; -x_319 = x_339; -x_320 = x_340; -x_321 = x_343; -x_322 = x_341; +x_317 = x_343; +x_318 = x_337; +x_319 = x_338; +x_320 = x_339; +x_321 = x_341; +x_322 = x_340; x_323 = x_342; x_324 = x_359; x_325 = x_344; @@ -22185,12 +22185,12 @@ else lean_object* x_360; lean_dec(x_348); x_360 = lean_box(0); -x_317 = x_337; -x_318 = x_338; -x_319 = x_339; -x_320 = x_340; -x_321 = x_343; -x_322 = x_341; +x_317 = x_343; +x_318 = x_337; +x_319 = x_338; +x_320 = x_339; +x_321 = x_341; +x_322 = x_340; x_323 = x_342; x_324 = x_360; x_325 = x_344; @@ -22208,8 +22208,8 @@ if (x_373 == 0) { lean_object* x_374; lean_object* x_375; lean_dec(x_368); +lean_dec(x_367); lean_dec_ref(x_366); -lean_dec(x_365); lean_dec(x_363); lean_dec(x_362); x_374 = l_Lean_Elab_Command_elabInitialize___closed__53; @@ -22234,8 +22234,8 @@ if (x_379 == 0) lean_object* x_380; lean_object* x_381; lean_dec(x_377); lean_dec(x_368); +lean_dec(x_367); lean_dec_ref(x_366); -lean_dec(x_365); lean_dec(x_363); lean_dec(x_362); x_380 = l_Lean_Elab_Command_elabInitialize___closed__53; @@ -22257,8 +22257,8 @@ if (x_384 == 0) lean_object* x_385; lean_object* x_386; lean_dec(x_382); lean_dec(x_368); +lean_dec(x_367); lean_dec_ref(x_366); -lean_dec(x_365); lean_dec(x_363); lean_dec(x_362); x_385 = l_Lean_Elab_Command_elabInitialize___closed__53; @@ -22277,8 +22277,8 @@ lean_ctor_set(x_388, 0, x_387); x_337 = x_362; x_338 = x_368; x_339 = x_363; -x_340 = x_365; -x_341 = x_366; +x_340 = x_366; +x_341 = x_365; x_342 = x_367; x_343 = x_388; x_344 = x_369; @@ -22296,8 +22296,8 @@ x_389 = lean_box(0); x_337 = x_362; x_338 = x_368; x_339 = x_363; -x_340 = x_365; -x_341 = x_366; +x_340 = x_366; +x_341 = x_365; x_342 = x_367; x_343 = x_389; x_344 = x_369; @@ -22322,8 +22322,8 @@ if (x_402 == 0) lean_object* x_403; lean_object* x_404; lean_dec(x_400); lean_dec(x_396); -lean_dec_ref(x_394); -lean_dec(x_393); +lean_dec(x_395); +lean_dec_ref(x_392); lean_dec(x_391); x_403 = l_Lean_Elab_Command_elabInitialize___closed__53; x_404 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_403, x_397, x_398); @@ -22337,8 +22337,8 @@ lean_object* x_405; lean_object* x_406; lean_object* x_407; uint8_t x_408; x_405 = l_Lean_Syntax_getArg(x_400, x_11); lean_dec(x_400); x_406 = l_Lean_Elab_Command_elabInitialize___closed__44; -lean_inc_ref(x_394); -x_407 = l_Lean_Name_mkStr4(x_5, x_6, x_394, x_406); +lean_inc_ref(x_392); +x_407 = l_Lean_Name_mkStr4(x_5, x_6, x_392, x_406); lean_inc(x_405); x_408 = l_Lean_Syntax_isOfKind(x_405, x_407); lean_dec(x_407); @@ -22347,8 +22347,8 @@ if (x_408 == 0) lean_object* x_409; lean_object* x_410; lean_dec(x_405); lean_dec(x_396); -lean_dec_ref(x_394); -lean_dec(x_393); +lean_dec(x_395); +lean_dec_ref(x_392); lean_dec(x_391); x_409 = l_Lean_Elab_Command_elabInitialize___closed__53; x_410 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_409, x_397, x_398); @@ -22365,11 +22365,11 @@ x_412 = l_Lean_Syntax_getArgs(x_411); lean_dec(x_411); x_413 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_413, 0, x_412); -x_362 = x_396; -x_363 = x_391; -x_364 = x_392; +x_362 = x_391; +x_363 = x_396; +x_364 = x_394; x_365 = x_393; -x_366 = x_394; +x_366 = x_392; x_367 = x_395; x_368 = x_413; x_369 = x_397; @@ -22384,11 +22384,11 @@ else lean_object* x_414; lean_dec(x_400); x_414 = lean_box(0); -x_362 = x_396; -x_363 = x_391; -x_364 = x_392; +x_362 = x_391; +x_363 = x_396; +x_364 = x_394; x_365 = x_393; -x_366 = x_394; +x_366 = x_392; x_367 = x_395; x_368 = x_414; x_369 = x_397; @@ -22402,8 +22402,8 @@ goto block_390; if (x_151 == 0) { lean_object* x_424; lean_object* x_425; +lean_dec(x_420); lean_dec_ref(x_419); -lean_dec(x_418); lean_dec(x_416); x_424 = l_Lean_Elab_Command_elabInitialize___closed__53; x_425 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_424, x_421, x_422); @@ -22425,8 +22425,8 @@ if (x_428 == 0) { lean_object* x_429; lean_object* x_430; lean_dec(x_426); +lean_dec(x_420); lean_dec_ref(x_419); -lean_dec(x_418); lean_dec(x_416); x_429 = l_Lean_Elab_Command_elabInitialize___closed__53; x_430 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_429, x_421, x_422); @@ -22446,8 +22446,8 @@ if (x_433 == 0) { lean_object* x_434; lean_object* x_435; lean_dec(x_431); +lean_dec(x_420); lean_dec_ref(x_419); -lean_dec(x_418); lean_dec(x_416); x_434 = l_Lean_Elab_Command_elabInitialize___closed__53; x_435 = l_Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__19___redArg(x_12, x_434, x_421, x_422); @@ -22461,9 +22461,9 @@ lean_object* x_436; x_436 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_436, 0, x_431); x_391 = x_416; -x_392 = x_417; +x_392 = x_419; x_393 = x_418; -x_394 = x_419; +x_394 = x_417; x_395 = x_420; x_396 = x_436; x_397 = x_421; @@ -22479,9 +22479,9 @@ lean_object* x_437; lean_dec(x_426); x_437 = lean_box(0); x_391 = x_416; -x_392 = x_417; +x_392 = x_419; x_393 = x_418; -x_394 = x_419; +x_394 = x_417; x_395 = x_420; x_396 = x_437; x_397 = x_421; @@ -22494,16 +22494,16 @@ goto block_415; block_473: { lean_object* x_463; lean_object* x_464; -lean_inc_ref(x_452); -x_463 = l_Array_append___redArg(x_452, x_462); +lean_inc_ref(x_440); +x_463 = l_Array_append___redArg(x_440, x_462); lean_dec_ref(x_462); -lean_inc(x_454); -lean_inc(x_459); +lean_inc(x_446); +lean_inc(x_445); x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_459); -lean_ctor_set(x_464, 1, x_454); +lean_ctor_set(x_464, 0, x_445); +lean_ctor_set(x_464, 1, x_446); lean_ctor_set(x_464, 2, x_463); -if (lean_obj_tag(x_443) == 0) +if (lean_obj_tag(x_456) == 0) { lean_object* x_465; x_465 = l_Lean_Elab_Command_elabInitialize___closed__41; @@ -22511,34 +22511,34 @@ x_80 = x_439; x_81 = x_440; x_82 = x_441; x_83 = x_442; -x_84 = x_444; -x_85 = lean_box(0); -x_86 = x_446; +x_84 = x_443; +x_85 = x_444; +x_86 = x_445; x_87 = x_447; -x_88 = x_448; -x_89 = x_449; -x_90 = x_450; -x_91 = x_452; -x_92 = x_451; -x_93 = x_453; -x_94 = x_454; -x_95 = x_464; -x_96 = x_455; -x_97 = x_457; -x_98 = x_456; +x_88 = x_446; +x_89 = x_448; +x_90 = x_449; +x_91 = x_464; +x_92 = x_450; +x_93 = x_451; +x_94 = x_452; +x_95 = x_454; +x_96 = x_453; +x_97 = lean_box(0); +x_98 = x_457; x_99 = x_458; x_100 = x_459; -x_101 = x_460; -x_102 = x_461; +x_101 = x_461; +x_102 = x_460; x_103 = x_465; goto block_150; } else { lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; -x_466 = lean_ctor_get(x_443, 0); +x_466 = lean_ctor_get(x_456, 0); lean_inc(x_466); -lean_dec_ref(x_443); +lean_dec_ref(x_456); x_467 = l_Lean_Elab_Command_elabInitialize___closed__42; x_468 = l_Lean_Elab_Command_elabInitialize___closed__43; x_469 = l_Lean_SourceInfo_fromRef(x_466, x_156); @@ -22546,32 +22546,32 @@ lean_dec(x_466); x_470 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_470, 0, x_469); lean_ctor_set(x_470, 1, x_467); -lean_inc(x_459); -x_471 = l_Lean_Syntax_node1(x_459, x_468, x_470); +lean_inc(x_445); +x_471 = l_Lean_Syntax_node1(x_445, x_468, x_470); x_472 = l_Array_mkArray1___redArg(x_471); x_80 = x_439; x_81 = x_440; x_82 = x_441; x_83 = x_442; -x_84 = x_444; -x_85 = lean_box(0); -x_86 = x_446; +x_84 = x_443; +x_85 = x_444; +x_86 = x_445; x_87 = x_447; -x_88 = x_448; -x_89 = x_449; -x_90 = x_450; -x_91 = x_452; -x_92 = x_451; -x_93 = x_453; -x_94 = x_454; -x_95 = x_464; -x_96 = x_455; -x_97 = x_457; -x_98 = x_456; +x_88 = x_446; +x_89 = x_448; +x_90 = x_449; +x_91 = x_464; +x_92 = x_450; +x_93 = x_451; +x_94 = x_452; +x_95 = x_454; +x_96 = x_453; +x_97 = lean_box(0); +x_98 = x_457; x_99 = x_458; x_100 = x_459; -x_101 = x_460; -x_102 = x_461; +x_101 = x_461; +x_102 = x_460; x_103 = x_472; goto block_150; } @@ -22579,52 +22579,52 @@ goto block_150; block_507: { lean_object* x_497; lean_object* x_498; -lean_inc_ref(x_485); -lean_inc(x_487); -lean_inc(x_492); +lean_inc_ref(x_474); +lean_inc(x_481); +lean_inc(x_480); x_497 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_497, 0, x_492); -lean_ctor_set(x_497, 1, x_487); -lean_ctor_set(x_497, 2, x_485); -lean_inc(x_487); -lean_inc(x_492); -x_498 = l_Lean_Syntax_node1(x_492, x_487, x_479); -if (lean_obj_tag(x_494) == 0) +lean_ctor_set(x_497, 0, x_480); +lean_ctor_set(x_497, 1, x_481); +lean_ctor_set(x_497, 2, x_474); +lean_inc(x_481); +lean_inc(x_480); +x_498 = l_Lean_Syntax_node1(x_480, x_481, x_476); +if (lean_obj_tag(x_491) == 0) { lean_object* x_499; x_499 = l_Lean_Elab_Command_elabInitialize___closed__41; x_439 = x_497; x_440 = x_474; x_441 = x_475; -x_442 = x_476; -x_443 = x_477; -x_444 = x_478; -x_445 = lean_box(0); -x_446 = x_480; -x_447 = x_481; -x_448 = x_482; -x_449 = x_483; -x_450 = x_498; -x_451 = x_484; -x_452 = x_485; -x_453 = x_486; +x_442 = x_477; +x_443 = x_478; +x_444 = x_479; +x_445 = x_480; +x_446 = x_481; +x_447 = x_482; +x_448 = x_483; +x_449 = x_484; +x_450 = x_485; +x_451 = x_498; +x_452 = x_486; +x_453 = x_488; x_454 = x_487; -x_455 = x_488; -x_456 = x_490; -x_457 = x_489; -x_458 = x_491; +x_455 = lean_box(0); +x_456 = x_489; +x_457 = x_495; +x_458 = x_490; x_459 = x_492; -x_460 = x_493; -x_461 = x_495; +x_460 = x_494; +x_461 = x_493; x_462 = x_499; goto block_473; } else { lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; -x_500 = lean_ctor_get(x_494, 0); +x_500 = lean_ctor_get(x_491, 0); lean_inc(x_500); -lean_dec_ref(x_494); +lean_dec_ref(x_491); x_501 = l___private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__12_spec__12___closed__16; x_502 = l_Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4___closed__7; x_503 = l_Lean_SourceInfo_fromRef(x_500, x_156); @@ -22632,32 +22632,32 @@ lean_dec(x_500); x_504 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_504, 0, x_503); lean_ctor_set(x_504, 1, x_501); -lean_inc(x_492); -x_505 = l_Lean_Syntax_node1(x_492, x_502, x_504); +lean_inc(x_480); +x_505 = l_Lean_Syntax_node1(x_480, x_502, x_504); x_506 = l_Array_mkArray1___redArg(x_505); x_439 = x_497; x_440 = x_474; x_441 = x_475; -x_442 = x_476; -x_443 = x_477; -x_444 = x_478; -x_445 = lean_box(0); -x_446 = x_480; -x_447 = x_481; -x_448 = x_482; -x_449 = x_483; -x_450 = x_498; -x_451 = x_484; -x_452 = x_485; -x_453 = x_486; +x_442 = x_477; +x_443 = x_478; +x_444 = x_479; +x_445 = x_480; +x_446 = x_481; +x_447 = x_482; +x_448 = x_483; +x_449 = x_484; +x_450 = x_485; +x_451 = x_498; +x_452 = x_486; +x_453 = x_488; x_454 = x_487; -x_455 = x_488; -x_456 = x_490; -x_457 = x_489; -x_458 = x_491; +x_455 = lean_box(0); +x_456 = x_489; +x_457 = x_495; +x_458 = x_490; x_459 = x_492; -x_460 = x_493; -x_461 = x_495; +x_460 = x_494; +x_461 = x_493; x_462 = x_506; goto block_473; } @@ -22665,52 +22665,52 @@ goto block_473; block_545: { lean_object* x_529; -x_529 = l_Lean_Elab_Command_getRef___redArg(x_524); +x_529 = l_Lean_Elab_Command_getRef___redArg(x_518); if (lean_obj_tag(x_529) == 0) { lean_object* x_530; lean_object* x_531; x_530 = lean_ctor_get(x_529, 0); lean_inc(x_530); lean_dec_ref(x_529); -x_531 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_524); +x_531 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_518); if (lean_obj_tag(x_531) == 0) { lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; x_532 = lean_ctor_get(x_531, 0); lean_inc(x_532); lean_dec_ref(x_531); -x_533 = lean_ctor_get(x_524, 5); +x_533 = lean_ctor_get(x_518, 5); x_534 = l_Lean_Parser_Command_visibility_ofBool(x_528); -x_535 = l_Lean_SourceInfo_fromRef(x_530, x_514); +x_535 = l_Lean_SourceInfo_fromRef(x_530, x_513); lean_dec(x_530); if (lean_obj_tag(x_533) == 0) { lean_object* x_536; lean_object* x_537; -x_536 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_512); +x_536 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_elabAttr___at___00Lean_Elab_elabAttrs___at___00Lean_Elab_elabDeclAttrs___at___00Lean_Elab_elabModifiers___at___00Lean_Elab_Command_elabDeclaration_spec__4_spec__4_spec__4_spec__9_spec__9_spec__22___redArg(x_520); x_537 = lean_ctor_get(x_536, 0); lean_inc(x_537); lean_dec_ref(x_536); x_474 = x_508; -x_475 = x_510; -x_476 = x_532; -x_477 = x_511; -x_478 = x_512; -x_479 = x_534; -x_480 = x_513; -x_481 = x_515; -x_482 = x_516; -x_483 = x_517; -x_484 = x_518; -x_485 = x_519; +x_475 = x_509; +x_476 = x_534; +x_477 = x_510; +x_478 = x_511; +x_479 = x_512; +x_480 = x_535; +x_481 = x_514; +x_482 = x_515; +x_483 = x_516; +x_484 = x_517; +x_485 = x_518; x_486 = x_520; x_487 = x_521; -x_488 = x_522; -x_489 = x_524; +x_488 = x_532; +x_489 = x_522; x_490 = x_523; x_491 = x_525; -x_492 = x_535; -x_493 = x_526; -x_494 = x_527; +x_492 = x_524; +x_493 = x_527; +x_494 = x_526; x_495 = x_537; x_496 = lean_box(0); goto block_507; @@ -22721,26 +22721,26 @@ lean_object* x_538; x_538 = lean_ctor_get(x_533, 0); lean_inc(x_538); x_474 = x_508; -x_475 = x_510; -x_476 = x_532; -x_477 = x_511; -x_478 = x_512; -x_479 = x_534; -x_480 = x_513; -x_481 = x_515; -x_482 = x_516; -x_483 = x_517; -x_484 = x_518; -x_485 = x_519; +x_475 = x_509; +x_476 = x_534; +x_477 = x_510; +x_478 = x_511; +x_479 = x_512; +x_480 = x_535; +x_481 = x_514; +x_482 = x_515; +x_483 = x_516; +x_484 = x_517; +x_485 = x_518; x_486 = x_520; x_487 = x_521; -x_488 = x_522; -x_489 = x_524; +x_488 = x_532; +x_489 = x_522; x_490 = x_523; x_491 = x_525; -x_492 = x_535; -x_493 = x_526; -x_494 = x_527; +x_492 = x_524; +x_493 = x_527; +x_494 = x_526; x_495 = x_538; x_496 = lean_box(0); goto block_507; @@ -22757,17 +22757,17 @@ lean_dec_ref(x_524); lean_dec(x_523); lean_dec(x_522); lean_dec(x_521); -lean_dec_ref(x_520); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec_ref(x_518); lean_dec(x_517); lean_dec(x_516); -lean_dec(x_515); -lean_dec_ref(x_513); +lean_dec_ref(x_515); +lean_dec(x_514); lean_dec(x_512); lean_dec(x_511); lean_dec(x_510); -lean_dec(x_508); +lean_dec_ref(x_509); +lean_dec_ref(x_508); x_539 = !lean_is_exclusive(x_531); if (x_539 == 0) { @@ -22795,17 +22795,17 @@ lean_dec_ref(x_524); lean_dec(x_523); lean_dec(x_522); lean_dec(x_521); -lean_dec_ref(x_520); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec_ref(x_518); lean_dec(x_517); lean_dec(x_516); -lean_dec(x_515); -lean_dec_ref(x_513); +lean_dec_ref(x_515); +lean_dec(x_514); lean_dec(x_512); lean_dec(x_511); lean_dec(x_510); -lean_dec(x_508); +lean_dec_ref(x_509); +lean_dec_ref(x_508); x_542 = !lean_is_exclusive(x_529); if (x_542 == 0) { @@ -22826,8 +22826,8 @@ return x_544; block_570: { lean_object* x_566; lean_object* x_567; lean_object* x_568; uint8_t x_569; -x_566 = l_Lean_Syntax_getArg(x_558, x_11); -x_567 = l_Lean_Syntax_getArg(x_558, x_153); +x_566 = l_Lean_Syntax_getArg(x_557, x_11); +x_567 = l_Lean_Syntax_getArg(x_557, x_153); lean_inc_ref(x_563); lean_inc(x_562); x_568 = l_Lean_Elab_addDeclarationRangesForBuiltin___at___00Lean_Elab_Command_elabInitialize_spec__1(x_562, x_566, x_567, x_563, x_564); @@ -22836,50 +22836,50 @@ x_569 = l_Lean_isPrivateName(x_562); if (x_569 == 0) { x_508 = x_546; -x_509 = lean_box(0); -x_510 = x_547; +x_509 = x_547; +x_510 = x_562; x_511 = x_548; -x_512 = x_564; -x_513 = x_549; -x_514 = x_550; -x_515 = x_551; -x_516 = x_552; -x_517 = x_562; -x_518 = x_553; -x_519 = x_554; -x_520 = x_555; -x_521 = x_556; -x_522 = x_557; -x_523 = x_558; -x_524 = x_563; -x_525 = x_559; -x_526 = x_560; -x_527 = x_561; +x_512 = x_549; +x_513 = x_550; +x_514 = x_551; +x_515 = x_552; +x_516 = x_553; +x_517 = x_554; +x_518 = x_563; +x_519 = lean_box(0); +x_520 = x_564; +x_521 = x_555; +x_522 = x_556; +x_523 = x_557; +x_524 = x_559; +x_525 = x_558; +x_526 = x_561; +x_527 = x_560; x_528 = x_156; goto block_545; } else { x_508 = x_546; -x_509 = lean_box(0); -x_510 = x_547; +x_509 = x_547; +x_510 = x_562; x_511 = x_548; -x_512 = x_564; -x_513 = x_549; -x_514 = x_550; -x_515 = x_551; -x_516 = x_552; -x_517 = x_562; -x_518 = x_553; -x_519 = x_554; -x_520 = x_555; -x_521 = x_556; -x_522 = x_557; -x_523 = x_558; -x_524 = x_563; -x_525 = x_559; -x_526 = x_560; -x_527 = x_561; +x_512 = x_549; +x_513 = x_550; +x_514 = x_551; +x_515 = x_552; +x_516 = x_553; +x_517 = x_554; +x_518 = x_563; +x_519 = lean_box(0); +x_520 = x_564; +x_521 = x_555; +x_522 = x_556; +x_523 = x_557; +x_524 = x_559; +x_525 = x_558; +x_526 = x_561; +x_527 = x_560; x_528 = x_550; goto block_545; } @@ -22888,222 +22888,222 @@ goto block_545; block_79: { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_31 = l_Array_append___redArg(x_16, x_30); +x_31 = l_Array_append___redArg(x_18, x_30); lean_dec_ref(x_30); lean_inc(x_23); -lean_inc(x_25); +lean_inc(x_14); x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_25); +lean_ctor_set(x_32, 0, x_14); lean_ctor_set(x_32, 1, x_23); lean_ctor_set(x_32, 2, x_31); -lean_inc_n(x_28, 2); -lean_inc(x_25); -x_33 = l_Lean_Syntax_node7(x_25, x_13, x_21, x_17, x_24, x_28, x_20, x_32, x_28); +lean_inc_n(x_19, 2); +lean_inc(x_14); +x_33 = l_Lean_Syntax_node7(x_14, x_13, x_27, x_17, x_16, x_19, x_28, x_32, x_19); x_34 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__20; x_35 = l_Lean_Elab_Command_elabInitialize___closed__4; -lean_inc(x_25); +lean_inc(x_14); x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_25); +lean_ctor_set(x_36, 0, x_14); lean_ctor_set(x_36, 1, x_35); x_37 = l_Lean_Elab_Command_elabInitialize___closed__6; x_38 = l_Lean_Elab_Command_elabInitialize___closed__8; x_39 = l_Lean_Elab_Command_elabInitialize___closed__9; -lean_inc(x_26); -lean_inc(x_14); -x_40 = l_Lean_addMacroScope(x_14, x_39, x_26); +lean_inc(x_21); +lean_inc(x_29); +x_40 = l_Lean_addMacroScope(x_29, x_39, x_21); x_41 = lean_box(0); -lean_inc(x_25); +lean_inc(x_14); x_42 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_42, 0, x_25); +lean_ctor_set(x_42, 0, x_14); lean_ctor_set(x_42, 1, x_38); lean_ctor_set(x_42, 2, x_40); lean_ctor_set(x_42, 3, x_41); -lean_inc(x_28); -lean_inc(x_25); -x_43 = l_Lean_Syntax_node2(x_25, x_37, x_42, x_28); +lean_inc(x_19); +lean_inc(x_14); +x_43 = l_Lean_Syntax_node2(x_14, x_37, x_42, x_19); x_44 = l_Lean_Elab_Command_elabInitialize___closed__11; x_45 = l_Lean_Elab_Command_elabInitialize___closed__12; -lean_inc_ref(x_27); -x_46 = l_Lean_Name_mkStr4(x_5, x_6, x_27, x_45); +lean_inc_ref(x_24); +x_46 = l_Lean_Name_mkStr4(x_5, x_6, x_24, x_45); x_47 = l_Lean_Elab_Command_elabInitialize___closed__13; -lean_inc(x_25); +lean_inc(x_14); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_25); +lean_ctor_set(x_48, 0, x_14); lean_ctor_set(x_48, 1, x_47); x_49 = l_Lean_Elab_Command_elabInitialize___closed__14; -lean_inc_ref(x_27); -x_50 = l_Lean_Name_mkStr4(x_5, x_6, x_27, x_49); +lean_inc_ref(x_24); +x_50 = l_Lean_Name_mkStr4(x_5, x_6, x_24, x_49); x_51 = l_Lean_Elab_Command_elabInitialize___closed__16; x_52 = l_Lean_Elab_Command_elabInitialize___closed__17; -lean_inc(x_26); -lean_inc(x_14); -x_53 = l_Lean_addMacroScope(x_14, x_52, x_26); +lean_inc(x_21); +lean_inc(x_29); +x_53 = l_Lean_addMacroScope(x_29, x_52, x_21); x_54 = l_Lean_Elab_Command_elabInitialize___closed__21; -lean_inc(x_25); +lean_inc(x_14); x_55 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_55, 0, x_25); +lean_ctor_set(x_55, 0, x_14); lean_ctor_set(x_55, 1, x_51); lean_ctor_set(x_55, 2, x_53); lean_ctor_set(x_55, 3, x_54); x_56 = l_Lean_Elab_Command_elabInitialize___closed__23; x_57 = l_Lean_Elab_Command_elabInitialize___closed__24; -x_58 = l_Lean_addMacroScope(x_14, x_57, x_26); +x_58 = l_Lean_addMacroScope(x_29, x_57, x_21); x_59 = l_Lean_Elab_Command_elabInitialize___closed__28; -lean_inc(x_25); +lean_inc(x_14); x_60 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_60, 0, x_25); +lean_ctor_set(x_60, 0, x_14); lean_ctor_set(x_60, 1, x_56); lean_ctor_set(x_60, 2, x_58); lean_ctor_set(x_60, 3, x_59); lean_inc(x_23); -lean_inc(x_25); -x_61 = l_Lean_Syntax_node1(x_25, x_23, x_60); -lean_inc(x_25); -x_62 = l_Lean_Syntax_node2(x_25, x_50, x_55, x_61); -lean_inc(x_25); -x_63 = l_Lean_Syntax_node2(x_25, x_46, x_48, x_62); -lean_inc(x_25); -x_64 = l_Lean_Syntax_node1(x_25, x_23, x_63); -lean_inc(x_28); -lean_inc(x_25); -x_65 = l_Lean_Syntax_node2(x_25, x_44, x_28, x_64); +lean_inc(x_14); +x_61 = l_Lean_Syntax_node1(x_14, x_23, x_60); +lean_inc(x_14); +x_62 = l_Lean_Syntax_node2(x_14, x_50, x_55, x_61); +lean_inc(x_14); +x_63 = l_Lean_Syntax_node2(x_14, x_46, x_48, x_62); +lean_inc(x_14); +x_64 = l_Lean_Syntax_node1(x_14, x_23, x_63); +lean_inc(x_19); +lean_inc(x_14); +x_65 = l_Lean_Syntax_node2(x_14, x_44, x_19, x_64); x_66 = l_Lean_Elab_Command_elabInitialize___closed__30; x_67 = l_Lean_Elab_Command_elabInitialize___closed__31; -lean_inc(x_25); +lean_inc(x_14); x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_25); +lean_ctor_set(x_68, 0, x_14); lean_ctor_set(x_68, 1, x_67); x_69 = l_Lean_Elab_Command_elabInitialize___closed__32; -x_70 = l_Lean_Name_mkStr4(x_5, x_6, x_27, x_69); -lean_inc(x_25); +x_70 = l_Lean_Name_mkStr4(x_5, x_6, x_24, x_69); +lean_inc(x_14); x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_25); +lean_ctor_set(x_71, 0, x_14); lean_ctor_set(x_71, 1, x_69); -lean_inc(x_25); -x_72 = l_Lean_Syntax_node2(x_25, x_70, x_71, x_29); +lean_inc(x_14); +x_72 = l_Lean_Syntax_node2(x_14, x_70, x_71, x_25); x_73 = l_Lean_Elab_Command_elabInitialize___closed__35; -lean_inc_n(x_28, 2); -lean_inc(x_25); -x_74 = l_Lean_Syntax_node2(x_25, x_73, x_28, x_28); -lean_inc(x_28); -lean_inc(x_25); -x_75 = l_Lean_Syntax_node4(x_25, x_66, x_68, x_72, x_74, x_28); -lean_inc(x_25); -x_76 = l_Lean_Syntax_node5(x_25, x_34, x_36, x_43, x_65, x_75, x_28); -x_77 = l_Lean_Syntax_node2(x_25, x_19, x_33, x_76); -x_78 = l_Lean_Elab_Command_elabCommand(x_77, x_15, x_18); +lean_inc_n(x_19, 2); +lean_inc(x_14); +x_74 = l_Lean_Syntax_node2(x_14, x_73, x_19, x_19); +lean_inc(x_19); +lean_inc(x_14); +x_75 = l_Lean_Syntax_node4(x_14, x_66, x_68, x_72, x_74, x_19); +lean_inc(x_14); +x_76 = l_Lean_Syntax_node5(x_14, x_34, x_36, x_43, x_65, x_75, x_19); +x_77 = l_Lean_Syntax_node2(x_14, x_26, x_33, x_76); +x_78 = l_Lean_Elab_Command_elabCommand(x_77, x_22, x_15); return x_78; } block_150: { lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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_104 = l_Array_append___redArg(x_91, x_103); +x_104 = l_Array_append___redArg(x_81, x_103); lean_dec_ref(x_103); -lean_inc(x_94); -lean_inc(x_100); +lean_inc(x_88); +lean_inc(x_86); x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_100); -lean_ctor_set(x_105, 1, x_94); +lean_ctor_set(x_105, 0, x_86); +lean_ctor_set(x_105, 1, x_88); lean_ctor_set(x_105, 2, x_104); lean_inc_n(x_80, 4); -lean_inc(x_100); -x_106 = l_Lean_Syntax_node7(x_100, x_13, x_80, x_80, x_90, x_80, x_95, x_105, x_80); +lean_inc(x_86); +x_106 = l_Lean_Syntax_node7(x_86, x_13, x_80, x_80, x_93, x_80, x_91, x_105, x_80); x_107 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDecl___closed__20; x_108 = l_Lean_Elab_Command_elabInitialize___closed__4; -lean_inc(x_100); +lean_inc(x_86); x_109 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_109, 0, x_100); +lean_ctor_set(x_109, 0, x_86); lean_ctor_set(x_109, 1, x_108); -lean_inc(x_83); -lean_inc(x_102); -x_110 = l_Lean_addMacroScope(x_102, x_87, x_83); -lean_inc(x_88); -lean_inc(x_100); +lean_inc(x_96); +lean_inc(x_98); +x_110 = l_Lean_addMacroScope(x_98, x_101, x_96); +lean_inc(x_85); +lean_inc(x_86); x_111 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_111, 0, x_100); -lean_ctor_set(x_111, 1, x_93); +lean_ctor_set(x_111, 0, x_86); +lean_ctor_set(x_111, 1, x_100); lean_ctor_set(x_111, 2, x_110); -lean_ctor_set(x_111, 3, x_88); +lean_ctor_set(x_111, 3, x_85); lean_inc(x_80); -lean_inc(x_100); -x_112 = l_Lean_Syntax_node2(x_100, x_101, x_111, x_80); +lean_inc(x_86); +x_112 = l_Lean_Syntax_node2(x_86, x_102, x_111, x_80); x_113 = l_Lean_Elab_Command_elabInitialize___closed__11; -lean_inc(x_100); +lean_inc(x_86); x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_100); -lean_ctor_set(x_114, 1, x_92); +lean_ctor_set(x_114, 0, x_86); +lean_ctor_set(x_114, 1, x_82); x_115 = l_Lean_Elab_Command_elabInitialize___closed__14; -lean_inc_ref(x_86); -x_116 = l_Lean_Name_mkStr4(x_5, x_6, x_86, x_115); +lean_inc_ref(x_87); +x_116 = l_Lean_Name_mkStr4(x_5, x_6, x_87, x_115); x_117 = l_Lean_Elab_Command_elabInitialize___closed__16; x_118 = l_Lean_Elab_Command_elabInitialize___closed__17; -x_119 = l_Lean_addMacroScope(x_102, x_118, x_83); +x_119 = l_Lean_addMacroScope(x_98, x_118, x_96); x_120 = l_Lean_Elab_Command_elabInitialize___closed__36; x_121 = l_Lean_Elab_Command_elabInitialize___closed__19; x_122 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_88); +lean_ctor_set(x_122, 1, x_85); x_123 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_123, 0, x_120); lean_ctor_set(x_123, 1, x_122); -lean_inc(x_100); +lean_inc(x_86); x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_100); +lean_ctor_set(x_124, 0, x_86); lean_ctor_set(x_124, 1, x_117); lean_ctor_set(x_124, 2, x_119); lean_ctor_set(x_124, 3, x_123); -lean_inc(x_94); -lean_inc(x_100); -x_125 = l_Lean_Syntax_node1(x_100, x_94, x_96); -lean_inc(x_100); -x_126 = l_Lean_Syntax_node2(x_100, x_116, x_124, x_125); -lean_inc(x_100); -x_127 = l_Lean_Syntax_node2(x_100, x_81, x_114, x_126); -lean_inc(x_94); -lean_inc(x_100); -x_128 = l_Lean_Syntax_node1(x_100, x_94, x_127); +lean_inc(x_88); +lean_inc(x_86); +x_125 = l_Lean_Syntax_node1(x_86, x_88, x_90); +lean_inc(x_86); +x_126 = l_Lean_Syntax_node2(x_86, x_116, x_124, x_125); +lean_inc(x_86); +x_127 = l_Lean_Syntax_node2(x_86, x_84, x_114, x_126); +lean_inc(x_88); +lean_inc(x_86); +x_128 = l_Lean_Syntax_node1(x_86, x_88, x_127); lean_inc(x_80); -lean_inc(x_100); -x_129 = l_Lean_Syntax_node2(x_100, x_113, x_80, x_128); +lean_inc(x_86); +x_129 = l_Lean_Syntax_node2(x_86, x_113, x_80, x_128); x_130 = l_Lean_Elab_Command_elabInitialize___closed__30; x_131 = l_Lean_Elab_Command_elabInitialize___closed__31; -lean_inc(x_100); +lean_inc(x_86); x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_100); +lean_ctor_set(x_132, 0, x_86); lean_ctor_set(x_132, 1, x_131); x_133 = l_Lean_Elab_Command_elabInitialize___closed__37; -lean_inc_ref(x_86); -x_134 = l_Lean_Name_mkStr4(x_5, x_6, x_86, x_133); +lean_inc_ref(x_87); +x_134 = l_Lean_Name_mkStr4(x_5, x_6, x_87, x_133); x_135 = l_Lean_Elab_Command_elabInitialize___closed__38; -lean_inc(x_100); +lean_inc(x_86); x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_100); +lean_ctor_set(x_136, 0, x_86); lean_ctor_set(x_136, 1, x_135); -x_137 = lean_mk_syntax_ident(x_89); +x_137 = lean_mk_syntax_ident(x_83); x_138 = l_Lean_Elab_Command_elabInitialize___closed__32; -x_139 = l_Lean_Name_mkStr4(x_5, x_6, x_86, x_138); -lean_inc(x_100); +x_139 = l_Lean_Name_mkStr4(x_5, x_6, x_87, x_138); +lean_inc(x_86); x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_100); +lean_ctor_set(x_140, 0, x_86); lean_ctor_set(x_140, 1, x_138); -lean_inc(x_100); -x_141 = l_Lean_Syntax_node2(x_100, x_139, x_140, x_99); +lean_inc(x_86); +x_141 = l_Lean_Syntax_node2(x_86, x_139, x_140, x_89); lean_inc(x_80); -lean_inc(x_100); -x_142 = l_Lean_Syntax_node4(x_100, x_134, x_136, x_80, x_137, x_141); +lean_inc(x_86); +x_142 = l_Lean_Syntax_node4(x_86, x_134, x_136, x_80, x_137, x_141); x_143 = l_Lean_Elab_Command_elabInitialize___closed__35; lean_inc_n(x_80, 2); -lean_inc(x_100); -x_144 = l_Lean_Syntax_node2(x_100, x_143, x_80, x_80); +lean_inc(x_86); +x_144 = l_Lean_Syntax_node2(x_86, x_143, x_80, x_80); lean_inc(x_80); -lean_inc(x_100); -x_145 = l_Lean_Syntax_node4(x_100, x_130, x_132, x_142, x_144, x_80); -lean_inc(x_100); -x_146 = l_Lean_Syntax_node5(x_100, x_107, x_109, x_112, x_129, x_145, x_80); -lean_inc(x_100); -x_147 = l_Lean_Syntax_node2(x_100, x_82, x_106, x_146); -x_148 = l_Lean_Syntax_node2(x_100, x_94, x_147, x_98); -x_149 = l_Lean_Elab_Command_elabCommand(x_148, x_97, x_84); +lean_inc(x_86); +x_145 = l_Lean_Syntax_node4(x_86, x_130, x_132, x_142, x_144, x_80); +lean_inc(x_86); +x_146 = l_Lean_Syntax_node5(x_86, x_107, x_109, x_112, x_129, x_145, x_80); +lean_inc(x_86); +x_147 = l_Lean_Syntax_node2(x_86, x_95, x_106, x_146); +x_148 = l_Lean_Syntax_node2(x_86, x_88, x_147, x_99); +x_149 = l_Lean_Elab_Command_elabCommand(x_148, x_92, x_94); return x_149; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index ecda71e07144..9a3fc979d2a5 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -130,7 +130,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at___00Lean_throwError___at lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkAuxFunction___closed__5; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew___closed__19; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__17___redArg(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_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkAuxFunction___closed__12; LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__15___boxed(lean_object**); @@ -271,7 +270,6 @@ LEAN_EXPORT lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__pri static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__0; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumCmd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_isEnumType___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstance_spec__1_spec__1_spec__1_spec__1_spec__1_spec__3_spec__3_spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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*); LEAN_EXPORT lean_object* l___private_Init_Data_Fin_Fold_0__Fin_foldlM_loop___at___00Array_ofFnM___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew_spec__4_spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatch(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*); @@ -4097,12 +4095,12 @@ goto block_24; block_180: { uint8_t x_171; -x_171 = lean_nat_dec_lt(x_169, x_170); +x_171 = lean_nat_dec_lt(x_168, x_170); if (x_171 == 0) { lean_dec(x_170); -lean_dec(x_169); -lean_dec_ref(x_168); +lean_dec_ref(x_169); +lean_dec(x_168); x_61 = x_60; x_62 = lean_box(0); goto block_167; @@ -4110,16 +4108,16 @@ goto block_167; else { size_t x_172; size_t x_173; lean_object* x_174; -x_172 = lean_usize_of_nat(x_169); -lean_dec(x_169); +x_172 = lean_usize_of_nat(x_168); +lean_dec(x_168); x_173 = lean_usize_of_nat(x_170); lean_dec(x_170); lean_inc(x_15); lean_inc_ref(x_14); lean_inc(x_13); lean_inc_ref(x_12); -x_174 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__14___redArg(x_41, x_168, x_172, x_173, x_12, x_13, x_14, x_15); -lean_dec_ref(x_168); +x_174 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__14___redArg(x_41, x_169, x_172, x_173, x_12, x_13, x_14, x_15); +lean_dec_ref(x_169); if (lean_obj_tag(x_174) == 0) { lean_object* x_175; uint8_t x_176; @@ -4202,16 +4200,16 @@ x_189 = lean_nat_dec_le(x_186, x_188); if (x_189 == 0) { lean_dec(x_186); -x_168 = x_184; -x_169 = x_185; +x_168 = x_185; +x_169 = x_184; x_170 = x_188; goto block_180; } else { lean_dec(x_188); -x_168 = x_184; -x_169 = x_185; +x_168 = x_185; +x_169 = x_184; x_170 = x_186; goto block_180; } @@ -5604,704 +5602,49 @@ return x_2; } } static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___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_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__3; -x_2 = l_Lean_Name_mkStr1(x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__3; -x_2 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__3; -x_3 = l_Lean_Name_mkStr2(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__6; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__7; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___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, 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_109; -lean_inc(x_15); -lean_inc_ref(x_14); -x_109 = l_Lean_Core_betaReduce(x_9, x_14, x_15); -if (lean_obj_tag(x_109) == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_126; lean_object* x_127; lean_object* x_155; uint8_t x_156; -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -lean_dec_ref(x_109); -x_111 = lean_ctor_get(x_1, 1); -lean_inc(x_111); -x_112 = lean_ctor_get(x_1, 2); -x_155 = lean_unsigned_to_nat(0u); -x_156 = lean_nat_dec_lt(x_155, x_112); -if (x_156 == 0) -{ -lean_inc_ref(x_5); -x_126 = x_5; -x_127 = lean_box(0); -goto block_154; -} -else -{ -lean_object* x_157; lean_object* x_158; -lean_inc_ref(x_5); -x_157 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt_spec__1___redArg(x_112, x_5, x_155, x_14); -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -lean_dec_ref(x_157); -x_126 = x_158; -x_127 = lean_box(0); -goto block_154; -} -block_125: -{ -lean_object* x_118; uint8_t x_119; -x_118 = lean_unsigned_to_nat(0u); -x_119 = lean_nat_dec_lt(x_118, x_111); -if (x_119 == 0) -{ -lean_dec(x_111); -x_17 = x_113; -x_18 = x_116; -x_19 = x_114; -x_20 = x_115; -x_21 = lean_box(0); -goto block_108; -} -else -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_114); -lean_ctor_set(x_120, 1, x_115); -x_121 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__13___redArg(x_111, x_120, x_118, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_111); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -lean_dec_ref(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(x_122); -x_17 = x_113; -x_18 = x_116; -x_19 = x_123; -x_20 = x_124; -x_21 = lean_box(0); -goto block_108; -} -} -block_154: -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_128 = lean_ctor_get(x_14, 5); -x_129 = lean_ctor_get(x_14, 10); -x_130 = lean_ctor_get(x_14, 11); -x_131 = lean_unsigned_to_nat(0u); -x_132 = 0; -x_133 = l_Lean_SourceInfo_fromRef(x_128, x_132); -x_134 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__4; -x_135 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__5; -lean_inc(x_130); -lean_inc(x_129); -x_136 = l_Lean_addMacroScope(x_129, x_135, x_130); -x_137 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__8; -x_138 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_138, 0, x_133); -lean_ctor_set(x_138, 1, x_134); -lean_ctor_set(x_138, 2, x_136); -lean_ctor_set(x_138, 3, x_137); -x_139 = lean_nat_dec_lt(x_131, x_4); -if (x_139 == 0) -{ -lean_dec(x_110); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_1); -lean_inc_ref(x_5); -x_113 = x_126; -x_114 = x_5; -x_115 = x_5; -x_116 = x_138; -x_117 = lean_box(0); -goto block_125; -} -else -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_140 = lean_box(x_139); -x_141 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_141, 0, x_138); -lean_ctor_set(x_141, 1, x_140); -lean_inc_ref(x_5); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_5); -lean_ctor_set(x_142, 1, x_141); -x_143 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_143, 0, x_5); -lean_ctor_set(x_143, 1, x_142); -lean_inc(x_15); -lean_inc_ref(x_14); -lean_inc(x_13); -lean_inc_ref(x_12); -x_144 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__15___redArg(x_1, x_4, x_6, x_8, x_110, x_7, x_4, x_131, x_143, x_131, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_110); -if (lean_obj_tag(x_144) == 0) -{ -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -lean_dec_ref(x_144); -x_146 = lean_ctor_get(x_145, 1); -lean_inc(x_146); -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -x_148 = lean_ctor_get(x_145, 0); -lean_inc(x_148); -lean_dec(x_145); -x_149 = lean_ctor_get(x_146, 0); -lean_inc(x_149); -lean_dec(x_146); -x_150 = lean_ctor_get(x_147, 0); -lean_inc(x_150); -lean_dec(x_147); -x_113 = x_126; -x_114 = x_148; -x_115 = x_149; -x_116 = x_150; -x_117 = lean_box(0); -goto block_125; -} -else -{ -uint8_t x_151; -lean_dec_ref(x_126); -lean_dec(x_111); -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_10); -lean_dec(x_3); -lean_dec_ref(x_2); -x_151 = !lean_is_exclusive(x_144); -if (x_151 == 0) -{ -return x_144; -} -else -{ -lean_object* x_152; lean_object* x_153; -x_152 = lean_ctor_get(x_144, 0); -lean_inc(x_152); -lean_dec(x_144); -x_153 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_153, 0, x_152); -return x_153; -} -} -} -} -} -else -{ -uint8_t x_159; -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_10); -lean_dec_ref(x_8); -lean_dec(x_7); -lean_dec_ref(x_6); -lean_dec_ref(x_5); -lean_dec(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_159 = !lean_is_exclusive(x_109); -if (x_159 == 0) -{ -return x_109; -} -else -{ -lean_object* x_160; lean_object* x_161; -x_160 = lean_ctor_get(x_109, 0); -lean_inc(x_160); -lean_dec(x_109); -x_161 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_161, 0, x_160); -return x_161; -} -} -block_108: -{ -lean_object* x_22; -lean_inc_ref(x_2); -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_22 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); -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_ref(x_2); -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_7(x_2, 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_25; lean_object* x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec_ref(x_24); -x_26 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); -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; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__15_spec__15___redArg___closed__51; -x_30 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__1; -x_31 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__2; -lean_inc(x_23); -x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_23); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_mk_syntax_ident(x_3); -lean_inc(x_33); -lean_inc(x_23); -x_34 = l_Lean_Syntax_node2(x_23, x_30, x_32, x_33); -x_35 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__11; -x_36 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__12; -x_37 = l_Array_reverse___redArg(x_19); -x_38 = l_Array_append___redArg(x_36, x_37); -lean_dec_ref(x_37); -lean_inc(x_23); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_23); -lean_ctor_set(x_39, 1, x_35); -lean_ctor_set(x_39, 2, x_38); -x_40 = l_Lean_Syntax_node2(x_23, x_29, x_34, x_39); -x_41 = lean_array_push(x_17, x_40); -lean_inc(x_25); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_25); -lean_ctor_set(x_42, 1, x_31); -lean_inc(x_25); -x_43 = l_Lean_Syntax_node2(x_25, x_30, x_42, x_33); -x_44 = l_Array_reverse___redArg(x_20); -x_45 = l_Array_append___redArg(x_36, x_44); -lean_dec_ref(x_44); -lean_inc(x_25); -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_25); -lean_ctor_set(x_46, 1, x_35); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Lean_Syntax_node2(x_25, x_29, x_43, x_46); -x_48 = lean_array_push(x_41, x_47); -x_49 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__8; -x_50 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__9; -lean_inc(x_28); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_28); -lean_ctor_set(x_51, 1, x_50); -x_52 = lean_array_size(x_48); -x_53 = 0; -x_54 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt_spec__0(x_52, x_53, x_48); -x_55 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__14; -x_56 = l_Lean_mkSepArray(x_54, x_55); -lean_dec_ref(x_54); -x_57 = l_Array_append___redArg(x_36, x_56); -lean_dec_ref(x_56); -lean_inc(x_28); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_28); -lean_ctor_set(x_58, 1, x_35); -lean_ctor_set(x_58, 2, x_57); -lean_inc(x_28); -x_59 = l_Lean_Syntax_node1(x_28, x_35, x_58); -x_60 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__15; -lean_inc(x_28); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_28); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Syntax_node4(x_28, x_49, x_51, x_59, x_61, x_18); -lean_ctor_set(x_26, 0, x_62); -return x_26; -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; 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; size_t x_87; size_t x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_63 = lean_ctor_get(x_26, 0); -lean_inc(x_63); -lean_dec(x_26); -x_64 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__15_spec__15___redArg___closed__51; -x_65 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__1; -x_66 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__2; -lean_inc(x_23); -x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_23); -lean_ctor_set(x_67, 1, x_66); -x_68 = lean_mk_syntax_ident(x_3); -lean_inc(x_68); -lean_inc(x_23); -x_69 = l_Lean_Syntax_node2(x_23, x_65, x_67, x_68); -x_70 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__11; -x_71 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__12; -x_72 = l_Array_reverse___redArg(x_19); -x_73 = l_Array_append___redArg(x_71, x_72); -lean_dec_ref(x_72); -lean_inc(x_23); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_23); -lean_ctor_set(x_74, 1, x_70); -lean_ctor_set(x_74, 2, x_73); -x_75 = l_Lean_Syntax_node2(x_23, x_64, x_69, x_74); -x_76 = lean_array_push(x_17, x_75); -lean_inc(x_25); -x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_25); -lean_ctor_set(x_77, 1, x_66); -lean_inc(x_25); -x_78 = l_Lean_Syntax_node2(x_25, x_65, x_77, x_68); -x_79 = l_Array_reverse___redArg(x_20); -x_80 = l_Array_append___redArg(x_71, x_79); -lean_dec_ref(x_79); -lean_inc(x_25); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_25); -lean_ctor_set(x_81, 1, x_70); -lean_ctor_set(x_81, 2, x_80); -x_82 = l_Lean_Syntax_node2(x_25, x_64, x_78, x_81); -x_83 = lean_array_push(x_76, x_82); -x_84 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__8; -x_85 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__9; -lean_inc(x_63); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_63); -lean_ctor_set(x_86, 1, x_85); -x_87 = lean_array_size(x_83); -x_88 = 0; -x_89 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt_spec__0(x_87, x_88, x_83); -x_90 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__14; -x_91 = l_Lean_mkSepArray(x_89, x_90); -lean_dec_ref(x_89); -x_92 = l_Array_append___redArg(x_71, x_91); -lean_dec_ref(x_91); -lean_inc(x_63); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_63); -lean_ctor_set(x_93, 1, x_70); -lean_ctor_set(x_93, 2, x_92); -lean_inc(x_63); -x_94 = l_Lean_Syntax_node1(x_63, x_70, x_93); -x_95 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__15; -lean_inc(x_63); -x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_63); -lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Syntax_node4(x_63, x_84, x_86, x_94, x_96, x_18); -x_98 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_98, 0, x_97); -return x_98; -} -} -else -{ -uint8_t x_99; -lean_dec(x_25); -lean_dec(x_23); -lean_dec_ref(x_20); -lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_17); -lean_dec(x_3); -x_99 = !lean_is_exclusive(x_26); -if (x_99 == 0) -{ -return x_26; -} -else -{ -lean_object* x_100; lean_object* x_101; -x_100 = lean_ctor_get(x_26, 0); -lean_inc(x_100); -lean_dec(x_26); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_100); -return x_101; -} -} -} -else -{ -uint8_t x_102; -lean_dec(x_23); -lean_dec_ref(x_20); -lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_17); -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_10); -lean_dec(x_3); -lean_dec_ref(x_2); -x_102 = !lean_is_exclusive(x_24); -if (x_102 == 0) -{ -return x_24; -} -else -{ -lean_object* x_103; lean_object* x_104; -x_103 = lean_ctor_get(x_24, 0); -lean_inc(x_103); -lean_dec(x_24); -x_104 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_104, 0, x_103); -return x_104; -} -} -} -else -{ -uint8_t x_105; -lean_dec_ref(x_20); -lean_dec_ref(x_19); -lean_dec(x_18); -lean_dec_ref(x_17); -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_10); -lean_dec(x_3); -lean_dec_ref(x_2); -x_105 = !lean_is_exclusive(x_22); -if (x_105 == 0) -{ -return x_22; -} -else -{ -lean_object* x_106; lean_object* x_107; -x_106 = lean_ctor_get(x_22, 0); -lean_inc(x_106); -lean_dec(x_22); -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_106); -return x_107; -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___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) { -_start: -{ -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_14; -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -lean_dec_ref(x_7); -lean_dec(x_4); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_6); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_5, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 1); -lean_inc(x_16); -lean_dec_ref(x_5); -lean_inc_ref(x_11); -lean_inc_ref(x_7); -lean_inc(x_15); -x_17 = l_Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__0(x_15, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec_ref(x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc_ref(x_19); -x_20 = lean_ctor_get(x_18, 4); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_19, 2); -lean_inc_ref(x_21); -lean_dec_ref(x_19); -x_22 = lean_alloc_closure((void*)(l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__13___redArg___lam__0___boxed), 7, 0); -lean_inc(x_4); -lean_inc_ref(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_1); -x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___boxed), 16, 7); -lean_closure_set(x_23, 0, x_1); -lean_closure_set(x_23, 1, x_22); -lean_closure_set(x_23, 2, x_15); -lean_closure_set(x_23, 3, x_20); -lean_closure_set(x_23, 4, x_2); -lean_closure_set(x_23, 5, x_3); -lean_closure_set(x_23, 6, x_4); -x_24 = 0; -lean_inc(x_12); -lean_inc_ref(x_11); -lean_inc(x_10); -lean_inc_ref(x_9); -lean_inc(x_8); -lean_inc_ref(x_7); -x_25 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__17___redArg(x_21, x_23, x_24, x_24, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec_ref(x_25); -x_27 = lean_array_push(x_6, x_26); -x_5 = x_16; -x_6 = x_27; -goto _start; -} -else -{ -uint8_t x_29; -lean_dec(x_16); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -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_ref(x_1); -x_29 = !lean_is_exclusive(x_25); -if (x_29 == 0) -{ -return x_25; -} -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_25, 0); -lean_inc(x_30); -lean_dec(x_25); -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_16); -lean_dec(x_15); -lean_dec(x_12); -lean_dec_ref(x_11); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec(x_8); -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_ref(x_1); -x_32 = !lean_is_exclusive(x_17); -if (x_32 == 0) -{ -return x_17; -} -else +_start: { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_17, 0); -lean_inc(x_33); -lean_dec(x_17); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -return x_34; +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_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__3; +x_2 = l_Lean_Name_mkStr1(x_1); +return x_2; } } +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__3; +x_2 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkElseAlt___closed__3; +x_3 = l_Lean_Name_mkStr2(x_2, x_1); +return x_3; +} } +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +static lean_object* _init_l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__8() { _start: { -lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg(x_1, x_2, x_3, x_4, x_6, x_7, x_9, x_10, x_11, x_12, x_13, x_14); -return x_16; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___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, 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_109; @@ -6346,8 +5689,8 @@ x_119 = lean_nat_dec_lt(x_118, x_111); if (x_119 == 0) { lean_dec(x_111); -x_17 = x_116; -x_18 = x_113; +x_17 = x_113; +x_18 = x_116; x_19 = x_114; x_20 = x_115; x_21 = lean_box(0); @@ -6369,8 +5712,8 @@ lean_inc(x_123); x_124 = lean_ctor_get(x_122, 1); lean_inc(x_124); lean_dec(x_122); -x_17 = x_116; -x_18 = x_113; +x_17 = x_113; +x_18 = x_116; x_19 = x_123; x_20 = x_124; x_21 = lean_box(0); @@ -6585,7 +5928,7 @@ lean_ctor_set(x_39, 0, x_23); lean_ctor_set(x_39, 1, x_35); lean_ctor_set(x_39, 2, x_38); x_40 = l_Lean_Syntax_node2(x_23, x_29, x_34, x_39); -x_41 = lean_array_push(x_18, x_40); +x_41 = lean_array_push(x_17, x_40); lean_inc(x_25); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_25); @@ -6628,7 +5971,7 @@ lean_inc(x_28); x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_28); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Syntax_node4(x_28, x_49, x_51, x_59, x_61, x_17); +x_62 = l_Lean_Syntax_node4(x_28, x_49, x_51, x_59, x_61, x_18); lean_ctor_set(x_26, 0, x_62); return x_26; } @@ -6660,7 +6003,7 @@ lean_ctor_set(x_74, 0, x_23); lean_ctor_set(x_74, 1, x_70); lean_ctor_set(x_74, 2, x_73); x_75 = l_Lean_Syntax_node2(x_23, x_64, x_69, x_74); -x_76 = lean_array_push(x_18, x_75); +x_76 = lean_array_push(x_17, x_75); lean_inc(x_25); x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_25); @@ -6703,7 +6046,7 @@ lean_inc(x_63); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_63); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Syntax_node4(x_63, x_84, x_86, x_94, x_96, x_17); +x_97 = l_Lean_Syntax_node4(x_63, x_84, x_86, x_94, x_96, x_18); x_98 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_98, 0, x_97); return x_98; @@ -6716,8 +6059,8 @@ lean_dec(x_25); lean_dec(x_23); lean_dec_ref(x_20); lean_dec_ref(x_19); -lean_dec_ref(x_18); -lean_dec(x_17); +lean_dec(x_18); +lean_dec_ref(x_17); lean_dec(x_3); x_99 = !lean_is_exclusive(x_26); if (x_99 == 0) @@ -6742,8 +6085,8 @@ uint8_t x_102; lean_dec(x_23); lean_dec_ref(x_20); lean_dec_ref(x_19); -lean_dec_ref(x_18); -lean_dec(x_17); +lean_dec(x_18); +lean_dec_ref(x_17); lean_dec(x_15); lean_dec_ref(x_14); lean_dec(x_13); @@ -6774,8 +6117,8 @@ else uint8_t x_105; lean_dec_ref(x_20); lean_dec_ref(x_19); -lean_dec_ref(x_18); -lean_dec(x_17); +lean_dec(x_18); +lean_dec_ref(x_17); lean_dec(x_15); lean_dec_ref(x_14); lean_dec(x_13); @@ -6803,6 +6146,159 @@ return x_107; } } } +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___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) { +_start: +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_14; +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +lean_dec_ref(x_7); +lean_dec(x_4); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_6); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_5, 1); +lean_inc(x_16); +lean_dec_ref(x_5); +lean_inc_ref(x_11); +lean_inc_ref(x_7); +lean_inc(x_15); +x_17 = l_Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__0(x_15, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec_ref(x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc_ref(x_19); +x_20 = lean_ctor_get(x_18, 4); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 2); +lean_inc_ref(x_21); +lean_dec_ref(x_19); +x_22 = lean_alloc_closure((void*)(l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__13___redArg___lam__0___boxed), 7, 0); +lean_inc(x_4); +lean_inc_ref(x_3); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___boxed), 16, 7); +lean_closure_set(x_23, 0, x_1); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_15); +lean_closure_set(x_23, 3, x_20); +lean_closure_set(x_23, 4, x_2); +lean_closure_set(x_23, 5, x_3); +lean_closure_set(x_23, 6, x_4); +x_24 = 0; +lean_inc(x_12); +lean_inc_ref(x_11); +lean_inc(x_10); +lean_inc_ref(x_9); +lean_inc(x_8); +lean_inc_ref(x_7); +x_25 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__17___redArg(x_21, x_23, x_24, x_24, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = lean_array_push(x_6, x_26); +x_5 = x_16; +x_6 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +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_ref(x_1); +x_29 = !lean_is_exclusive(x_25); +if (x_29 == 0) +{ +return x_25; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_25, 0); +lean_inc(x_30); +lean_dec(x_25); +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_16); +lean_dec(x_15); +lean_dec(x_12); +lean_dec_ref(x_11); +lean_dec(x_10); +lean_dec_ref(x_9); +lean_dec(x_8); +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_ref(x_1); +x_32 = !lean_is_exclusive(x_17); +if (x_32 == 0) +{ +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_17, 0); +lean_inc(x_33); +lean_dec(x_17); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_16; +x_16 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg(x_1, x_2, x_3, x_4, x_6, x_7, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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: { @@ -6854,7 +6350,7 @@ lean_inc(x_3); lean_inc_ref(x_2); lean_inc_ref(x_4); lean_inc_ref(x_1); -x_24 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___redArg___lam__1___boxed), 16, 7); +x_24 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18_spec__18___redArg___lam__1___boxed), 16, 7); lean_closure_set(x_24, 0, x_1); lean_closure_set(x_24, 1, x_23); lean_closure_set(x_24, 2, x_16); @@ -7530,15 +7026,6 @@ lean_dec(x_5); return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -lean_object* x_17; -x_17 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_spec__18___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, x_13, x_14, x_15); -lean_dec(x_4); -return x_17; -} -} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchOld_mkAlts_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_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { @@ -8471,12 +7958,12 @@ goto block_25; block_180: { uint8_t x_171; -x_171 = lean_nat_dec_lt(x_168, x_170); +x_171 = lean_nat_dec_lt(x_169, x_170); if (x_171 == 0) { lean_dec(x_170); -lean_dec_ref(x_169); -lean_dec(x_168); +lean_dec(x_169); +lean_dec_ref(x_168); lean_dec(x_41); x_61 = x_60; x_62 = lean_box(0); @@ -8485,16 +7972,16 @@ goto block_167; else { size_t x_172; size_t x_173; lean_object* x_174; -x_172 = lean_usize_of_nat(x_168); -lean_dec(x_168); +x_172 = lean_usize_of_nat(x_169); +lean_dec(x_169); x_173 = lean_usize_of_nat(x_170); lean_dec(x_170); lean_inc(x_16); lean_inc_ref(x_15); lean_inc(x_14); lean_inc_ref(x_13); -x_174 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew_spec__1___redArg(x_41, x_169, x_172, x_173, x_13, x_14, x_15, x_16); -lean_dec_ref(x_169); +x_174 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew_spec__1___redArg(x_41, x_168, x_172, x_173, x_13, x_14, x_15, x_16); +lean_dec_ref(x_168); lean_dec(x_41); if (lean_obj_tag(x_174) == 0) { @@ -8576,16 +8063,16 @@ x_189 = lean_nat_dec_le(x_186, x_188); if (x_189 == 0) { lean_dec(x_186); -x_168 = x_185; -x_169 = x_184; +x_168 = x_184; +x_169 = x_185; x_170 = x_188; goto block_180; } else { lean_dec(x_188); -x_168 = x_185; -x_169 = x_184; +x_168 = x_184; +x_169 = x_185; x_170 = x_186; goto block_180; } @@ -9399,12 +8886,12 @@ goto block_26; block_181: { uint8_t x_172; -x_172 = lean_nat_dec_lt(x_169, x_171); +x_172 = lean_nat_dec_lt(x_170, x_171); if (x_172 == 0) { lean_dec(x_171); -lean_dec_ref(x_170); -lean_dec(x_169); +lean_dec(x_170); +lean_dec_ref(x_169); lean_dec(x_42); x_62 = x_61; x_63 = lean_box(0); @@ -9413,16 +8900,16 @@ goto block_168; else { size_t x_173; size_t x_174; lean_object* x_175; -x_173 = lean_usize_of_nat(x_169); -lean_dec(x_169); +x_173 = lean_usize_of_nat(x_170); +lean_dec(x_170); x_174 = lean_usize_of_nat(x_171); lean_dec(x_171); lean_inc(x_17); lean_inc_ref(x_16); lean_inc(x_15); lean_inc_ref(x_14); -x_175 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew_spec__1___redArg(x_42, x_170, x_173, x_174, x_14, x_15, x_16, x_17); -lean_dec_ref(x_170); +x_175 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkMatchNew_spec__1___redArg(x_42, x_169, x_173, x_174, x_14, x_15, x_16, x_17); +lean_dec_ref(x_169); lean_dec(x_42); if (lean_obj_tag(x_175) == 0) { @@ -9503,16 +8990,16 @@ x_190 = lean_nat_dec_le(x_187, x_189); if (x_190 == 0) { lean_dec(x_187); -x_169 = x_186; -x_170 = x_185; +x_169 = x_185; +x_170 = x_186; x_171 = x_189; goto block_181; } else { lean_dec(x_189); -x_169 = x_186; -x_170 = x_185; +x_169 = x_185; +x_170 = x_186; x_171 = x_187; goto block_181; } @@ -16645,8 +16132,8 @@ x_50 = l_Lean_getMainModule___at___00__private_Lean_Elab_Deriving_BEq_0__Lean_El x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); lean_dec_ref(x_50); -x_6 = x_47; -x_7 = x_43; +x_6 = x_43; +x_7 = x_47; x_8 = x_49; x_9 = x_51; x_10 = lean_box(0); @@ -16657,8 +16144,8 @@ else lean_object* x_52; x_52 = lean_ctor_get(x_48, 0); lean_inc(x_52); -x_6 = x_47; -x_7 = x_43; +x_6 = x_43; +x_7 = x_47; x_8 = x_49; x_9 = x_52; x_10 = lean_box(0); @@ -16780,7 +16267,7 @@ x_21 = l_Lean_Syntax_node1(x_8, x_18, x_20); x_22 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstance___lam__1___closed__9; x_23 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstance___lam__1___closed__11; x_24 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstance___lam__1___closed__12; -x_25 = l_Lean_addMacroScope(x_9, x_24, x_6); +x_25 = l_Lean_addMacroScope(x_9, x_24, x_7); x_26 = lean_box(0); lean_inc(x_8); x_27 = lean_alloc_ctor(3, 4, 0); @@ -16799,7 +16286,7 @@ lean_inc(x_8); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_8); lean_ctor_set(x_32, 1, x_31); -x_33 = lean_mk_syntax_ident(x_7); +x_33 = lean_mk_syntax_ident(x_6); lean_inc(x_8); x_34 = l_Lean_Syntax_node1(x_8, x_16, x_33); x_35 = l_Lean_Syntax_node5(x_8, x_12, x_13, x_15, x_30, x_32, x_34); diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index 8a71620ed451..7f5ae7e1453d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -74,7 +74,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_ static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__0___redArg___closed__11; lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___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_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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_object*, lean_object*); static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__40; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__6; lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -225,7 +224,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_m static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__2___closed__2; lean_object* l_Nat_reprFast(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__0___redArg___closed__12; static lean_object* l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__0___redArg___closed__6; @@ -4609,7 +4607,7 @@ goto block_119; block_62: { lean_object* x_27; -x_27 = lean_apply_10(x_2, x_3, x_22, x_26, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +x_27 = lean_apply_10(x_2, x_3, x_24, x_26, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; @@ -4619,34 +4617,34 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_29 = lean_ctor_get(x_27, 0); x_30 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__0; -x_31 = l_Lean_Name_mkStr4(x_25, x_21, x_17, x_30); +x_31 = l_Lean_Name_mkStr4(x_23, x_22, x_25, x_30); x_32 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__1; -lean_inc(x_23); +lean_inc(x_18); x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_23); +lean_ctor_set(x_33, 0, x_18); lean_ctor_set(x_33, 1, x_32); -x_34 = lean_array_size(x_18); +x_34 = lean_array_size(x_19); x_35 = 0; -x_36 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_34, x_35, x_18); +x_36 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_34, x_35, x_19); x_37 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__2; x_38 = l_Lean_mkSepArray(x_36, x_37); lean_dec_ref(x_36); -x_39 = l_Array_append___redArg(x_19, x_38); +x_39 = l_Array_append___redArg(x_17, x_38); lean_dec_ref(x_38); -lean_inc(x_20); -lean_inc(x_23); +lean_inc(x_21); +lean_inc(x_18); x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_23); -lean_ctor_set(x_40, 1, x_20); +lean_ctor_set(x_40, 0, x_18); +lean_ctor_set(x_40, 1, x_21); lean_ctor_set(x_40, 2, x_39); -lean_inc(x_23); -x_41 = l_Lean_Syntax_node1(x_23, x_20, x_40); +lean_inc(x_18); +x_41 = l_Lean_Syntax_node1(x_18, x_21, x_40); x_42 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__3; -lean_inc(x_23); +lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 0, x_18); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Syntax_node4(x_23, x_31, x_33, x_41, x_43, x_29); +x_44 = l_Lean_Syntax_node4(x_18, x_31, x_33, x_41, x_43, x_29); lean_ctor_set(x_27, 0, x_44); return x_27; } @@ -4657,34 +4655,34 @@ x_45 = lean_ctor_get(x_27, 0); lean_inc(x_45); lean_dec(x_27); x_46 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__0; -x_47 = l_Lean_Name_mkStr4(x_25, x_21, x_17, x_46); +x_47 = l_Lean_Name_mkStr4(x_23, x_22, x_25, x_46); x_48 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__1; -lean_inc(x_23); +lean_inc(x_18); x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_23); +lean_ctor_set(x_49, 0, x_18); lean_ctor_set(x_49, 1, x_48); -x_50 = lean_array_size(x_18); +x_50 = lean_array_size(x_19); x_51 = 0; -x_52 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_50, x_51, x_18); +x_52 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_50, x_51, x_19); x_53 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__2; x_54 = l_Lean_mkSepArray(x_52, x_53); lean_dec_ref(x_52); -x_55 = l_Array_append___redArg(x_19, x_54); +x_55 = l_Array_append___redArg(x_17, x_54); lean_dec_ref(x_54); -lean_inc(x_20); -lean_inc(x_23); +lean_inc(x_21); +lean_inc(x_18); x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_23); -lean_ctor_set(x_56, 1, x_20); +lean_ctor_set(x_56, 0, x_18); +lean_ctor_set(x_56, 1, x_21); lean_ctor_set(x_56, 2, x_55); -lean_inc(x_23); -x_57 = l_Lean_Syntax_node1(x_23, x_20, x_56); +lean_inc(x_18); +x_57 = l_Lean_Syntax_node1(x_18, x_21, x_56); x_58 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__3; -lean_inc(x_23); +lean_inc(x_18); x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_23); +lean_ctor_set(x_59, 0, x_18); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Syntax_node4(x_23, x_47, x_49, x_57, x_59, x_45); +x_60 = l_Lean_Syntax_node4(x_18, x_47, x_49, x_57, x_59, x_45); x_61 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_61, 0, x_60); return x_61; @@ -4693,11 +4691,11 @@ return x_61; else { lean_dec_ref(x_25); -lean_dec(x_23); -lean_dec_ref(x_21); -lean_dec(x_20); +lean_dec_ref(x_23); +lean_dec_ref(x_22); +lean_dec(x_21); lean_dec_ref(x_19); -lean_dec_ref(x_18); +lean_dec(x_18); lean_dec_ref(x_17); return x_27; } @@ -4743,15 +4741,15 @@ if (x_88 == 0) lean_object* x_89; lean_dec_ref(x_66); x_89 = lean_box(0); -x_17 = x_73; -x_18 = x_85; -x_19 = x_81; -x_20 = x_80; -x_21 = x_72; -x_22 = x_64; -x_23 = x_70; -x_24 = lean_box(0); -x_25 = x_71; +x_17 = x_81; +x_18 = x_70; +x_19 = x_85; +x_20 = lean_box(0); +x_21 = x_80; +x_22 = x_72; +x_23 = x_71; +x_24 = x_64; +x_25 = x_73; x_26 = x_89; goto block_62; } @@ -4760,15 +4758,15 @@ else lean_object* x_90; x_90 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_90, 0, x_66); -x_17 = x_73; -x_18 = x_85; -x_19 = x_81; -x_20 = x_80; -x_21 = x_72; -x_22 = x_64; -x_23 = x_70; -x_24 = lean_box(0); -x_25 = x_71; +x_17 = x_81; +x_18 = x_70; +x_19 = x_85; +x_20 = lean_box(0); +x_21 = x_80; +x_22 = x_72; +x_23 = x_71; +x_24 = x_64; +x_25 = x_73; x_26 = x_90; goto block_62; } @@ -5064,7 +5062,7 @@ goto block_119; block_62: { lean_object* x_27; -x_27 = lean_apply_10(x_2, x_3, x_19, x_26, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +x_27 = lean_apply_10(x_2, x_3, x_24, x_26, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; @@ -5074,34 +5072,34 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_29 = lean_ctor_get(x_27, 0); x_30 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__0; -x_31 = l_Lean_Name_mkStr4(x_17, x_21, x_23, x_30); +x_31 = l_Lean_Name_mkStr4(x_25, x_21, x_22, x_30); x_32 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__1; -lean_inc(x_25); +lean_inc(x_23); x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 0, x_23); lean_ctor_set(x_33, 1, x_32); -x_34 = lean_array_size(x_20); +x_34 = lean_array_size(x_18); x_35 = 0; -x_36 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_34, x_35, x_20); +x_36 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_34, x_35, x_18); x_37 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__2; x_38 = l_Lean_mkSepArray(x_36, x_37); lean_dec_ref(x_36); -x_39 = l_Array_append___redArg(x_18, x_38); +x_39 = l_Array_append___redArg(x_17, x_38); lean_dec_ref(x_38); -lean_inc(x_24); -lean_inc(x_25); +lean_inc(x_19); +lean_inc(x_23); x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_24); +lean_ctor_set(x_40, 0, x_23); +lean_ctor_set(x_40, 1, x_19); lean_ctor_set(x_40, 2, x_39); -lean_inc(x_25); -x_41 = l_Lean_Syntax_node1(x_25, x_24, x_40); +lean_inc(x_23); +x_41 = l_Lean_Syntax_node1(x_23, x_19, x_40); x_42 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__3; -lean_inc(x_25); +lean_inc(x_23); x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_25); +lean_ctor_set(x_43, 0, x_23); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Syntax_node4(x_25, x_31, x_33, x_41, x_43, x_29); +x_44 = l_Lean_Syntax_node4(x_23, x_31, x_33, x_41, x_43, x_29); lean_ctor_set(x_27, 0, x_44); return x_27; } @@ -5112,34 +5110,34 @@ x_45 = lean_ctor_get(x_27, 0); lean_inc(x_45); lean_dec(x_27); x_46 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__0; -x_47 = l_Lean_Name_mkStr4(x_17, x_21, x_23, x_46); +x_47 = l_Lean_Name_mkStr4(x_25, x_21, x_22, x_46); x_48 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__1; -lean_inc(x_25); +lean_inc(x_23); x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_25); +lean_ctor_set(x_49, 0, x_23); lean_ctor_set(x_49, 1, x_48); -x_50 = lean_array_size(x_20); +x_50 = lean_array_size(x_18); x_51 = 0; -x_52 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_50, x_51, x_20); +x_52 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_50, x_51, x_18); x_53 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__2; x_54 = l_Lean_mkSepArray(x_52, x_53); lean_dec_ref(x_52); -x_55 = l_Array_append___redArg(x_18, x_54); +x_55 = l_Array_append___redArg(x_17, x_54); lean_dec_ref(x_54); -lean_inc(x_24); -lean_inc(x_25); +lean_inc(x_19); +lean_inc(x_23); x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_25); -lean_ctor_set(x_56, 1, x_24); +lean_ctor_set(x_56, 0, x_23); +lean_ctor_set(x_56, 1, x_19); lean_ctor_set(x_56, 2, x_55); -lean_inc(x_25); -x_57 = l_Lean_Syntax_node1(x_25, x_24, x_56); +lean_inc(x_23); +x_57 = l_Lean_Syntax_node1(x_23, x_19, x_56); x_58 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__3; -lean_inc(x_25); +lean_inc(x_23); x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_25); +lean_ctor_set(x_59, 0, x_23); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Syntax_node4(x_25, x_47, x_49, x_57, x_59, x_45); +x_60 = l_Lean_Syntax_node4(x_23, x_47, x_49, x_57, x_59, x_45); x_61 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_61, 0, x_60); return x_61; @@ -5147,11 +5145,11 @@ return x_61; } else { -lean_dec(x_25); -lean_dec(x_24); -lean_dec_ref(x_23); +lean_dec_ref(x_25); +lean_dec(x_23); +lean_dec_ref(x_22); lean_dec_ref(x_21); -lean_dec_ref(x_20); +lean_dec(x_19); lean_dec_ref(x_18); lean_dec_ref(x_17); return x_27; @@ -5198,15 +5196,15 @@ if (x_88 == 0) lean_object* x_89; lean_dec_ref(x_66); x_89 = lean_box(0); -x_17 = x_71; -x_18 = x_81; -x_19 = x_64; -x_20 = x_85; +x_17 = x_81; +x_18 = x_85; +x_19 = x_80; +x_20 = lean_box(0); x_21 = x_72; -x_22 = lean_box(0); -x_23 = x_73; -x_24 = x_80; -x_25 = x_70; +x_22 = x_73; +x_23 = x_70; +x_24 = x_64; +x_25 = x_71; x_26 = x_89; goto block_62; } @@ -5215,15 +5213,15 @@ else lean_object* x_90; x_90 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_90, 0, x_66); -x_17 = x_71; -x_18 = x_81; -x_19 = x_64; -x_20 = x_85; +x_17 = x_81; +x_18 = x_85; +x_19 = x_80; +x_20 = lean_box(0); x_21 = x_72; -x_22 = lean_box(0); -x_23 = x_73; -x_24 = x_80; -x_25 = x_70; +x_22 = x_73; +x_23 = x_70; +x_24 = x_64; +x_25 = x_71; x_26 = x_90; goto block_62; } @@ -11359,8 +11357,8 @@ x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_25); lean_ctor_set(x_89, 1, x_87); x_90 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__3; -x_91 = l_Array_zip___redArg(x_17, x_18); -lean_dec_ref(x_18); +x_91 = l_Array_zip___redArg(x_18, x_17); +lean_dec_ref(x_17); x_92 = lean_array_size(x_91); lean_inc_ref(x_84); lean_inc(x_25); @@ -11375,8 +11373,8 @@ x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_25); lean_ctor_set(x_98, 1, x_97); x_99 = lean_mk_syntax_ident(x_1); -x_100 = lean_array_size(x_17); -x_101 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__3___redArg(x_100, x_16, x_17); +x_100 = lean_array_size(x_18); +x_101 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__3___redArg(x_100, x_16, x_18); x_102 = l_Array_append___redArg(x_83, x_101); lean_dec_ref(x_101); lean_inc(x_25); @@ -11465,8 +11463,8 @@ lean_ctor_set(x_144, 1, x_140); lean_ctor_set(x_144, 2, x_142); lean_ctor_set(x_144, 3, x_143); x_16 = x_132; -x_17 = x_133; -x_18 = x_130; +x_17 = x_130; +x_18 = x_133; x_19 = x_144; x_20 = x_136; x_21 = x_137; @@ -11532,8 +11530,8 @@ lean_inc(x_149); x_172 = l_Lean_Syntax_node1(x_149, x_156, x_171); x_173 = l_Lean_Syntax_node2(x_149, x_150, x_155, x_172); x_16 = x_132; -x_17 = x_133; -x_18 = x_130; +x_17 = x_130; +x_18 = x_133; x_19 = x_173; x_20 = x_145; x_21 = x_146; @@ -11695,393 +11693,6 @@ x_16 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lea return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_175; uint8_t x_176; -x_120 = lean_unsigned_to_nat(0u); -x_175 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__7; -x_176 = lean_nat_dec_lt(x_120, x_2); -if (x_176 == 0) -{ -lean_dec(x_14); -lean_dec_ref(x_11); -lean_dec_ref(x_6); -x_121 = x_175; -x_122 = x_175; -x_123 = lean_box(0); -goto block_174; -} -else -{ -lean_object* x_177; lean_object* x_178; -x_177 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__27; -lean_inc_ref(x_13); -x_178 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__5___redArg(x_5, x_6, x_7, x_2, x_177, x_120, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_178) == 0) -{ -lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -lean_dec_ref(x_178); -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); -lean_dec(x_179); -x_121 = x_180; -x_122 = x_181; -x_123 = lean_box(0); -goto block_174; -} -else -{ -uint8_t x_182; -lean_dec_ref(x_13); -lean_dec_ref(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_182 = !lean_is_exclusive(x_178); -if (x_182 == 0) -{ -return x_178; -} -else -{ -lean_object* x_183; lean_object* x_184; -x_183 = lean_ctor_get(x_178, 0); -lean_inc(x_183); -lean_dec(x_178); -x_184 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_184, 0, x_183); -return x_184; -} -} -} -block_119: -{ -uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; size_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; 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; lean_object* x_117; lean_object* x_118; -x_24 = 0; -x_25 = l_Lean_SourceInfo_fromRef(x_20, x_24); -lean_dec(x_20); -x_26 = l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader___closed__0; -x_27 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__2; -x_28 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__3; -x_29 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__31; -x_30 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__39; -x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__41; -x_32 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__6; -x_33 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__6; -lean_inc(x_25); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_25); -lean_ctor_set(x_34, 1, x_33); -x_35 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__8; -x_36 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__10; -x_37 = lean_box(0); -lean_inc(x_22); -lean_inc(x_21); -x_38 = l_Lean_addMacroScope(x_21, x_37, x_22); -x_39 = lean_box(0); -x_40 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__28; -lean_inc(x_25); -x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_36); -lean_ctor_set(x_41, 2, x_38); -lean_ctor_set(x_41, 3, x_40); -lean_inc(x_25); -x_42 = l_Lean_Syntax_node1(x_25, x_35, x_41); -lean_inc(x_25); -x_43 = l_Lean_Syntax_node2(x_25, x_32, x_34, x_42); -x_44 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__1; -x_45 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__3; -lean_inc(x_22); -lean_inc(x_21); -x_46 = l_Lean_addMacroScope(x_21, x_45, x_22); -x_47 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__6; -lean_inc(x_25); -x_48 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_48, 0, x_25); -lean_ctor_set(x_48, 1, x_44); -lean_ctor_set(x_48, 2, x_46); -lean_ctor_set(x_48, 3, x_47); -x_49 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__8; -x_50 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__10; -x_51 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__11; -lean_inc(x_22); -lean_inc(x_21); -x_52 = l_Lean_addMacroScope(x_21, x_51, x_22); -x_53 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__15; -lean_inc(x_25); -x_54 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_54, 0, x_25); -lean_ctor_set(x_54, 1, x_50); -lean_ctor_set(x_54, 2, x_52); -lean_ctor_set(x_54, 3, x_53); -lean_inc(x_1); -x_55 = lean_erase_macro_scopes(x_1); -x_56 = l_Lean_Name_getString_x21(x_55); -lean_dec(x_55); -x_57 = lean_box(2); -x_58 = l_Lean_Syntax_mkStrLit(x_56, x_57); -lean_inc(x_2); -x_59 = l_Nat_reprFast(x_2); -x_60 = l_Lean_Syntax_mkNumLit(x_59, x_57); -lean_inc(x_25); -x_61 = l_Lean_Syntax_node4(x_25, x_49, x_54, x_58, x_60, x_19); -lean_inc(x_25); -x_62 = l_Lean_Syntax_node2(x_25, x_29, x_48, x_61); -x_63 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; -lean_inc(x_25); -x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_25); -lean_ctor_set(x_64, 1, x_63); -lean_inc_ref(x_64); -lean_inc(x_43); -lean_inc(x_25); -x_65 = l_Lean_Syntax_node3(x_25, x_31, x_43, x_62, x_64); -x_66 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__42; -lean_inc(x_25); -x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_25); -lean_ctor_set(x_67, 1, x_66); -x_68 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__8; -x_69 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__9; -lean_inc(x_22); -lean_inc(x_21); -x_70 = l_Lean_addMacroScope(x_21, x_69, x_22); -x_71 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__13; -lean_inc(x_25); -x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_25); -lean_ctor_set(x_72, 1, x_68); -lean_ctor_set(x_72, 2, x_70); -lean_ctor_set(x_72, 3, x_71); -lean_inc(x_25); -x_73 = l_Lean_Syntax_node3(x_25, x_30, x_65, x_67, x_72); -x_74 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__0___redArg___closed__16; -x_75 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__0___redArg___closed__17; -lean_inc(x_25); -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_25); -lean_ctor_set(x_76, 1, x_74); -x_77 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct_spec__0___redArg___closed__19; -x_78 = l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__0___redArg___closed__9; -x_79 = l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__0___redArg___closed__10; -x_80 = l_Lean_addMacroScope(x_21, x_79, x_22); -lean_inc(x_25); -x_81 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_81, 0, x_25); -lean_ctor_set(x_81, 1, x_78); -lean_ctor_set(x_81, 2, x_80); -lean_ctor_set(x_81, 3, x_39); -lean_inc(x_25); -x_82 = l_Lean_Syntax_node1(x_25, x_49, x_81); -x_83 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__24; -lean_inc(x_25); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_25); -lean_ctor_set(x_84, 1, x_49); -lean_ctor_set(x_84, 2, x_83); -x_85 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__3; -lean_inc(x_25); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_25); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__0; -x_88 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; -lean_inc(x_25); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_25); -lean_ctor_set(x_89, 1, x_87); -x_90 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__3; -x_91 = l_Array_zip___redArg(x_18, x_16); -lean_dec_ref(x_16); -x_92 = lean_array_size(x_91); -lean_inc_ref(x_84); -lean_inc(x_25); -x_93 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__2(x_26, x_27, x_28, x_25, x_84, x_92, x_17, x_91); -x_94 = l_Array_append___redArg(x_83, x_93); -lean_dec_ref(x_93); -x_95 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; -x_96 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; -x_97 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; -lean_inc(x_25); -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_25); -lean_ctor_set(x_98, 1, x_97); -x_99 = lean_mk_syntax_ident(x_1); -x_100 = lean_array_size(x_18); -x_101 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__3___redArg(x_100, x_17, x_18); -x_102 = l_Array_append___redArg(x_83, x_101); -lean_dec_ref(x_101); -lean_inc(x_25); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_25); -lean_ctor_set(x_103, 1, x_49); -lean_ctor_set(x_103, 2, x_102); -lean_inc(x_25); -x_104 = l_Lean_Syntax_node2(x_25, x_29, x_99, x_103); -lean_inc(x_25); -x_105 = l_Lean_Syntax_node1(x_25, x_49, x_104); -lean_inc(x_25); -x_106 = l_Lean_Syntax_node2(x_25, x_96, x_98, x_105); -lean_inc_ref(x_84); -lean_inc(x_25); -x_107 = l_Lean_Syntax_node2(x_25, x_95, x_106, x_84); -x_108 = lean_array_push(x_94, x_107); -lean_inc(x_25); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_25); -lean_ctor_set(x_109, 1, x_49); -lean_ctor_set(x_109, 2, x_108); -lean_inc(x_25); -x_110 = l_Lean_Syntax_node1(x_25, x_90, x_109); -lean_inc(x_25); -x_111 = l_Lean_Syntax_node2(x_25, x_88, x_89, x_110); -lean_inc(x_25); -x_112 = l_Lean_Syntax_node4(x_25, x_77, x_82, x_84, x_86, x_111); -lean_inc(x_25); -x_113 = l_Lean_Syntax_node2(x_25, x_75, x_76, x_112); -lean_inc(x_25); -x_114 = l_Lean_Syntax_node3(x_25, x_31, x_43, x_113, x_64); -lean_inc(x_25); -x_115 = l_Lean_Syntax_node1(x_25, x_49, x_114); -x_116 = l_Lean_Syntax_node2(x_25, x_29, x_73, x_115); -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_2); -x_118 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_118, 0, x_117); -return x_118; -} -block_174: -{ -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; size_t x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; -x_124 = lean_ctor_get(x_3, 2); -x_125 = lean_array_get_borrowed(x_4, x_124, x_120); -lean_inc(x_125); -x_126 = lean_mk_syntax_ident(x_125); -x_127 = lean_array_get_size(x_121); -x_128 = lean_mk_empty_array_with_capacity(x_127); -lean_inc_ref(x_13); -lean_inc(x_127); -x_129 = l_Array_mapFinIdxM_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__0___redArg(x_5, x_126, x_121, x_127, x_120, x_128, x_13); -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -lean_dec_ref(x_129); -x_131 = lean_array_size(x_121); -x_132 = 0; -x_133 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__1(x_131, x_132, x_121); -x_134 = lean_array_get_size(x_122); -x_135 = lean_nat_dec_eq(x_127, x_134); -lean_dec(x_134); -lean_dec(x_127); -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_dec_ref(x_122); -x_136 = lean_ctor_get(x_13, 5); -lean_inc(x_136); -x_137 = lean_ctor_get(x_13, 10); -lean_inc(x_137); -x_138 = lean_ctor_get(x_13, 11); -lean_inc(x_138); -lean_dec_ref(x_13); -x_139 = l_Lean_SourceInfo_fromRef(x_136, x_135); -x_140 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__15; -x_141 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__16; -lean_inc(x_138); -lean_inc(x_137); -x_142 = l_Lean_addMacroScope(x_137, x_141, x_138); -x_143 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__20; -x_144 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_144, 0, x_139); -lean_ctor_set(x_144, 1, x_140); -lean_ctor_set(x_144, 2, x_142); -lean_ctor_set(x_144, 3, x_143); -x_16 = x_130; -x_17 = x_132; -x_18 = x_133; -x_19 = x_144; -x_20 = x_136; -x_21 = x_137; -x_22 = x_138; -x_23 = lean_box(0); -goto block_119; -} -else -{ -lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_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; size_t x_161; lean_object* x_162; size_t x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_145 = lean_ctor_get(x_13, 5); -lean_inc(x_145); -x_146 = lean_ctor_get(x_13, 10); -lean_inc(x_146); -x_147 = lean_ctor_get(x_13, 11); -lean_inc(x_147); -lean_dec_ref(x_13); -x_148 = 0; -x_149 = l_Lean_SourceInfo_fromRef(x_145, x_148); -x_150 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__31; -x_151 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__22; -x_152 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__23; -lean_inc(x_147); -lean_inc(x_146); -x_153 = l_Lean_addMacroScope(x_146, x_152, x_147); -x_154 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___closed__26; -lean_inc(x_149); -x_155 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_155, 0, x_149); -lean_ctor_set(x_155, 1, x_151); -lean_ctor_set(x_155, 2, x_153); -lean_ctor_set(x_155, 3, x_154); -x_156 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__8; -x_157 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lam__0___closed__15; -x_158 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lam__0___closed__16; -lean_inc(x_149); -x_159 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_159, 0, x_149); -lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__24; -x_161 = lean_array_size(x_122); -x_162 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__4(x_161, x_132, x_122); -x_163 = lean_array_size(x_162); -x_164 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__13(x_163, x_132, x_162); -x_165 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts_spec__17_spec__17___redArg___lam__0___closed__2; -x_166 = l_Lean_mkSepArray(x_164, x_165); -lean_dec_ref(x_164); -x_167 = l_Array_append___redArg(x_160, x_166); -lean_dec_ref(x_166); -lean_inc(x_149); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_149); -lean_ctor_set(x_168, 1, x_156); -lean_ctor_set(x_168, 2, x_167); -x_169 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct_spec__0___redArg___closed__43; -lean_inc(x_149); -x_170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_170, 0, x_149); -lean_ctor_set(x_170, 1, x_169); -lean_inc(x_149); -x_171 = l_Lean_Syntax_node3(x_149, x_157, x_159, x_168, x_170); -lean_inc(x_149); -x_172 = l_Lean_Syntax_node1(x_149, x_156, x_171); -x_173 = l_Lean_Syntax_node2(x_149, x_150, x_155, x_172); -x_16 = x_130; -x_17 = x_132; -x_18 = x_133; -x_19 = x_173; -x_20 = x_145; -x_21 = x_146; -x_22 = x_147; -x_23 = lean_box(0); -goto block_119; -} -} -} -} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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: { @@ -12132,7 +11743,7 @@ lean_inc_ref(x_2); lean_inc_ref(x_1); lean_inc(x_4); lean_inc_ref(x_3); -x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___redArg___lam__0___boxed), 15, 6); +x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6_spec__6___redArg___lam__0___boxed), 15, 6); lean_closure_set(x_23, 0, x_16); lean_closure_set(x_23, 1, x_21); lean_closure_set(x_23, 2, x_3); @@ -12620,20 +12231,6 @@ lean_dec(x_5); return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_spec__6___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_12, x_13, x_14); -lean_dec(x_12); -lean_dec(x_10); -lean_dec_ref(x_9); -lean_dec_ref(x_8); -lean_dec_ref(x_7); -lean_dec_ref(x_3); -return x_16; -} -} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts_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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index b07cac0a1e24..1d58b2a6d443 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -70,6 +70,7 @@ static lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_ LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_throwUnknownIdentifierAt___at___00Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_isEnumType___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstance_spec__0_spec__0_spec__0_spec__0_spec__0_spec__2_spec__2_spec__2___redArg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__0_spec__0_spec__0_spec__0___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_getConstInfo___at___00Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0_spec__0___redArg___closed__12; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isEnumType___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstance_spec__0___boxed(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_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__17; static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__7; @@ -558,6 +559,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord extern lean_object* l_Lean_instInhabitedInductiveVal_default; LEAN_EXPORT lean_object* l___private_Init_Data_Fin_Fold_0__Fin_foldlM_loop___at___00Array_ofFnM___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchNew_spec__2_spec__2___redArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___redArg___lam__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstantAt___at___00Lean_throwUnknownConstant___at___00Lean_getConstInfo___at___00Lean_isEnumType___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstance_spec__0_spec__0_spec__0_spec__0___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_isEnumType___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstance_spec__0_spec__0_spec__0_spec__0_spec__0(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___lam__0___closed__14; @@ -4090,6 +4092,1084 @@ if (x_240 == 0) { lean_dec(x_233); lean_dec_ref(x_6); +x_17 = x_234; +x_18 = x_236; +x_19 = x_237; +x_20 = x_5; +x_21 = lean_box(0); +goto block_231; +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_241 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_241, 0, x_237); +lean_ctor_set(x_241, 1, x_5); +x_242 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_242, 0, x_236); +lean_ctor_set(x_242, 1, x_241); +lean_inc(x_15); +lean_inc_ref(x_14); +lean_inc(x_13); +lean_inc_ref(x_12); +x_243 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg(x_235, x_6, x_8, x_233, x_4, x_242, x_239, x_12, x_13, x_14, x_15); +lean_dec(x_233); +if (lean_obj_tag(x_243) == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_244 = lean_ctor_get(x_243, 0); +lean_inc(x_244); +lean_dec_ref(x_243); +x_245 = lean_ctor_get(x_244, 1); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 0); +lean_inc(x_246); +lean_dec(x_244); +x_247 = lean_ctor_get(x_245, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_245, 1); +lean_inc(x_248); +lean_dec(x_245); +x_17 = x_234; +x_18 = x_246; +x_19 = x_247; +x_20 = x_248; +x_21 = lean_box(0); +goto block_231; +} +else +{ +uint8_t x_249; +lean_dec_ref(x_234); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_249 = !lean_is_exclusive(x_243); +if (x_249 == 0) +{ +return x_243; +} +else +{ +lean_object* x_250; lean_object* x_251; +x_250 = lean_ctor_get(x_243, 0); +lean_inc(x_250); +lean_dec(x_243); +x_251 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_251, 0, x_250); +return x_251; +} +} +} +} +block_264: +{ +lean_object* x_257; uint8_t x_258; +x_257 = lean_unsigned_to_nat(0u); +x_258 = lean_nat_dec_lt(x_257, x_253); +if (x_258 == 0) +{ +lean_inc_ref(x_7); +x_234 = x_255; +x_235 = x_253; +x_236 = x_7; +x_237 = x_7; +x_238 = lean_box(0); +goto block_252; +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_inc_ref(x_7); +x_259 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_259, 0, x_7); +lean_ctor_set(x_259, 1, x_7); +x_260 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__15___redArg(x_253, x_259, x_257, x_14); +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +lean_dec_ref(x_260); +x_262 = lean_ctor_get(x_261, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_261, 1); +lean_inc(x_263); +lean_dec(x_261); +x_234 = x_255; +x_235 = x_253; +x_236 = x_262; +x_237 = x_263; +x_238 = lean_box(0); +goto block_252; +} +} +} +else +{ +uint8_t x_269; +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_10); +lean_dec_ref(x_7); +lean_dec_ref(x_6); +lean_dec_ref(x_5); +lean_dec(x_3); +lean_dec_ref(x_2); +x_269 = !lean_is_exclusive(x_232); +if (x_269 == 0) +{ +return x_232; +} +else +{ +lean_object* x_270; lean_object* x_271; +x_270 = lean_ctor_get(x_232, 0); +lean_inc(x_270); +lean_dec(x_232); +x_271 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_271, 0, x_270); +return x_271; +} +} +block_231: +{ +lean_object* x_22; +lean_inc_ref(x_2); +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_22 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +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_ref(x_2); +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_7(x_2, 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_25; lean_object* x_26; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +lean_dec_ref(x_24); +lean_inc_ref(x_2); +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_26 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +lean_dec_ref(x_26); +lean_inc_ref(x_2); +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_28 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec_ref(x_28); +x_30 = lean_ctor_get(x_14, 5); +lean_inc(x_30); +x_31 = lean_ctor_get(x_14, 10); +lean_inc(x_31); +x_32 = lean_ctor_get(x_14, 11); +lean_inc(x_32); +x_33 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__1; +x_34 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__3; +lean_inc(x_32); +lean_inc(x_31); +x_35 = l_Lean_addMacroScope(x_31, x_34, x_32); +x_36 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__7; +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_29); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_35); +lean_ctor_set(x_37, 3, x_36); +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_38 = lean_apply_8(x_20, x_37, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +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_ref(x_2); +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_40 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +lean_dec_ref(x_40); +lean_inc_ref(x_2); +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_42 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +lean_dec_ref(x_42); +x_44 = lean_apply_7(x_2, x_10, x_11, x_12, x_13, x_14, x_15, lean_box(0)); +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; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; 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; size_t x_83; size_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; size_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; size_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; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_46 = lean_ctor_get(x_44, 0); +x_47 = 0; +x_48 = l_Lean_SourceInfo_fromRef(x_30, x_47); +lean_dec(x_30); +x_49 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___lam__0___closed__4; +x_50 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__9; +x_51 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__10; +lean_inc(x_48); +x_52 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_52, 0, x_48); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_mk_syntax_ident(x_3); +lean_inc(x_53); +lean_inc(x_48); +x_54 = l_Lean_Syntax_node2(x_48, x_50, x_52, x_53); +x_55 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___lam__0___closed__15; +x_56 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__11; +x_57 = l_Array_append___redArg(x_56, x_18); +lean_dec_ref(x_18); +lean_inc(x_48); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_48); +lean_ctor_set(x_58, 1, x_55); +lean_ctor_set(x_58, 2, x_57); +x_59 = l_Lean_Syntax_node2(x_48, x_49, x_54, x_58); +lean_inc(x_23); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_23); +lean_ctor_set(x_60, 1, x_51); +lean_inc(x_23); +x_61 = l_Lean_Syntax_node2(x_23, x_50, x_60, x_53); +x_62 = l_Array_append___redArg(x_56, x_19); +lean_dec_ref(x_19); +lean_inc(x_23); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_23); +lean_ctor_set(x_63, 1, x_55); +lean_ctor_set(x_63, 2, x_62); +x_64 = l_Lean_Syntax_node2(x_23, x_49, x_61, x_63); +x_65 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__12; +x_66 = lean_array_push(x_65, x_59); +lean_inc(x_64); +lean_inc_ref(x_66); +x_67 = lean_array_push(x_66, x_64); +lean_inc_ref(x_17); +x_68 = l_Array_append___redArg(x_17, x_67); +lean_dec_ref(x_67); +x_69 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___closed__8; +x_70 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___closed__9; +lean_inc(x_25); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_25); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_Syntax_node1(x_25, x_69, x_71); +x_73 = lean_array_push(x_66, x_72); +lean_inc_ref(x_17); +x_74 = l_Array_append___redArg(x_17, x_73); +lean_dec_ref(x_73); +lean_inc(x_27); +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_27); +lean_ctor_set(x_75, 1, x_70); +x_76 = l_Lean_Syntax_node1(x_27, x_69, x_75); +x_77 = lean_array_push(x_65, x_76); +x_78 = lean_array_push(x_77, x_64); +x_79 = l_Array_append___redArg(x_17, x_78); +lean_dec_ref(x_78); +x_80 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__14; +x_81 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__15; +lean_inc(x_41); +x_82 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_82, 0, x_41); +lean_ctor_set(x_82, 1, x_81); +x_83 = lean_array_size(x_68); +x_84 = 0; +x_85 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_83, x_84, x_68); +x_86 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__17; +x_87 = l_Lean_mkSepArray(x_85, x_86); +lean_dec_ref(x_85); +x_88 = l_Array_append___redArg(x_56, x_87); +lean_dec_ref(x_87); +lean_inc(x_41); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_41); +lean_ctor_set(x_89, 1, x_55); +lean_ctor_set(x_89, 2, x_88); +lean_inc(x_41); +x_90 = l_Lean_Syntax_node1(x_41, x_55, x_89); +x_91 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__18; +lean_inc(x_41); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_41); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Lean_Syntax_node4(x_41, x_80, x_82, x_90, x_92, x_39); +lean_inc(x_43); +x_94 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_94, 0, x_43); +lean_ctor_set(x_94, 1, x_81); +x_95 = lean_array_size(x_74); +x_96 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_95, x_84, x_74); +x_97 = l_Lean_mkSepArray(x_96, x_86); +lean_dec_ref(x_96); +x_98 = l_Array_append___redArg(x_56, x_97); +lean_dec_ref(x_97); +lean_inc(x_43); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_43); +lean_ctor_set(x_99, 1, x_55); +lean_ctor_set(x_99, 2, x_98); +lean_inc(x_43); +x_100 = l_Lean_Syntax_node1(x_43, x_55, x_99); +lean_inc(x_43); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_43); +lean_ctor_set(x_101, 1, x_91); +x_102 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__20; +x_103 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__22; +lean_inc(x_32); +lean_inc(x_31); +x_104 = l_Lean_addMacroScope(x_31, x_103, x_32); +x_105 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__26; +lean_inc(x_43); +x_106 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_106, 0, x_43); +lean_ctor_set(x_106, 1, x_102); +lean_ctor_set(x_106, 2, x_104); +lean_ctor_set(x_106, 3, x_105); +x_107 = l_Lean_Syntax_node4(x_43, x_80, x_94, x_100, x_101, x_106); +lean_inc(x_46); +x_108 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_108, 0, x_46); +lean_ctor_set(x_108, 1, x_81); +x_109 = lean_array_size(x_79); +x_110 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_109, x_84, x_79); +x_111 = l_Lean_mkSepArray(x_110, x_86); +lean_dec_ref(x_110); +x_112 = l_Array_append___redArg(x_56, x_111); +lean_dec_ref(x_111); +lean_inc(x_46); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_46); +lean_ctor_set(x_113, 1, x_55); +lean_ctor_set(x_113, 2, x_112); +lean_inc(x_46); +x_114 = l_Lean_Syntax_node1(x_46, x_55, x_113); +lean_inc(x_46); +x_115 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_115, 0, x_46); +lean_ctor_set(x_115, 1, x_91); +x_116 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__28; +x_117 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__30; +x_118 = l_Lean_addMacroScope(x_31, x_117, x_32); +x_119 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__34; +lean_inc(x_46); +x_120 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_120, 0, x_46); +lean_ctor_set(x_120, 1, x_116); +lean_ctor_set(x_120, 2, x_118); +lean_ctor_set(x_120, 3, x_119); +x_121 = l_Lean_Syntax_node4(x_46, x_80, x_108, x_114, x_115, x_120); +x_122 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__35; +x_123 = lean_array_push(x_122, x_93); +x_124 = lean_array_push(x_123, x_107); +x_125 = lean_array_push(x_124, x_121); +lean_ctor_set(x_44, 0, x_125); +return x_44; +} +else +{ +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; 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; size_t x_163; size_t x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; size_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; size_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_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; +x_126 = lean_ctor_get(x_44, 0); +lean_inc(x_126); +lean_dec(x_44); +x_127 = 0; +x_128 = l_Lean_SourceInfo_fromRef(x_30, x_127); +lean_dec(x_30); +x_129 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___lam__0___closed__4; +x_130 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__9; +x_131 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__10; +lean_inc(x_128); +x_132 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_132, 0, x_128); +lean_ctor_set(x_132, 1, x_131); +x_133 = lean_mk_syntax_ident(x_3); +lean_inc(x_133); +lean_inc(x_128); +x_134 = l_Lean_Syntax_node2(x_128, x_130, x_132, x_133); +x_135 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___lam__0___closed__15; +x_136 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__11; +x_137 = l_Array_append___redArg(x_136, x_18); +lean_dec_ref(x_18); +lean_inc(x_128); +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_128); +lean_ctor_set(x_138, 1, x_135); +lean_ctor_set(x_138, 2, x_137); +x_139 = l_Lean_Syntax_node2(x_128, x_129, x_134, x_138); +lean_inc(x_23); +x_140 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_140, 0, x_23); +lean_ctor_set(x_140, 1, x_131); +lean_inc(x_23); +x_141 = l_Lean_Syntax_node2(x_23, x_130, x_140, x_133); +x_142 = l_Array_append___redArg(x_136, x_19); +lean_dec_ref(x_19); +lean_inc(x_23); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_23); +lean_ctor_set(x_143, 1, x_135); +lean_ctor_set(x_143, 2, x_142); +x_144 = l_Lean_Syntax_node2(x_23, x_129, x_141, x_143); +x_145 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__12; +x_146 = lean_array_push(x_145, x_139); +lean_inc(x_144); +lean_inc_ref(x_146); +x_147 = lean_array_push(x_146, x_144); +lean_inc_ref(x_17); +x_148 = l_Array_append___redArg(x_17, x_147); +lean_dec_ref(x_147); +x_149 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___closed__8; +x_150 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__14___redArg___closed__9; +lean_inc(x_25); +x_151 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_151, 0, x_25); +lean_ctor_set(x_151, 1, x_150); +x_152 = l_Lean_Syntax_node1(x_25, x_149, x_151); +x_153 = lean_array_push(x_146, x_152); +lean_inc_ref(x_17); +x_154 = l_Array_append___redArg(x_17, x_153); +lean_dec_ref(x_153); +lean_inc(x_27); +x_155 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_155, 0, x_27); +lean_ctor_set(x_155, 1, x_150); +x_156 = l_Lean_Syntax_node1(x_27, x_149, x_155); +x_157 = lean_array_push(x_145, x_156); +x_158 = lean_array_push(x_157, x_144); +x_159 = l_Array_append___redArg(x_17, x_158); +lean_dec_ref(x_158); +x_160 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__14; +x_161 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__15; +lean_inc(x_41); +x_162 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_162, 0, x_41); +lean_ctor_set(x_162, 1, x_161); +x_163 = lean_array_size(x_148); +x_164 = 0; +x_165 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_163, x_164, x_148); +x_166 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__17; +x_167 = l_Lean_mkSepArray(x_165, x_166); +lean_dec_ref(x_165); +x_168 = l_Array_append___redArg(x_136, x_167); +lean_dec_ref(x_167); +lean_inc(x_41); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_41); +lean_ctor_set(x_169, 1, x_135); +lean_ctor_set(x_169, 2, x_168); +lean_inc(x_41); +x_170 = l_Lean_Syntax_node1(x_41, x_135, x_169); +x_171 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__18; +lean_inc(x_41); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_41); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Lean_Syntax_node4(x_41, x_160, x_162, x_170, x_172, x_39); +lean_inc(x_43); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_43); +lean_ctor_set(x_174, 1, x_161); +x_175 = lean_array_size(x_154); +x_176 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_175, x_164, x_154); +x_177 = l_Lean_mkSepArray(x_176, x_166); +lean_dec_ref(x_176); +x_178 = l_Array_append___redArg(x_136, x_177); +lean_dec_ref(x_177); +lean_inc(x_43); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_43); +lean_ctor_set(x_179, 1, x_135); +lean_ctor_set(x_179, 2, x_178); +lean_inc(x_43); +x_180 = l_Lean_Syntax_node1(x_43, x_135, x_179); +lean_inc(x_43); +x_181 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_181, 0, x_43); +lean_ctor_set(x_181, 1, x_171); +x_182 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__20; +x_183 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__22; +lean_inc(x_32); +lean_inc(x_31); +x_184 = l_Lean_addMacroScope(x_31, x_183, x_32); +x_185 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__26; +lean_inc(x_43); +x_186 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_186, 0, x_43); +lean_ctor_set(x_186, 1, x_182); +lean_ctor_set(x_186, 2, x_184); +lean_ctor_set(x_186, 3, x_185); +x_187 = l_Lean_Syntax_node4(x_43, x_160, x_174, x_180, x_181, x_186); +lean_inc(x_126); +x_188 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_188, 0, x_126); +lean_ctor_set(x_188, 1, x_161); +x_189 = lean_array_size(x_159); +x_190 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__13(x_189, x_164, x_159); +x_191 = l_Lean_mkSepArray(x_190, x_166); +lean_dec_ref(x_190); +x_192 = l_Array_append___redArg(x_136, x_191); +lean_dec_ref(x_191); +lean_inc(x_126); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_126); +lean_ctor_set(x_193, 1, x_135); +lean_ctor_set(x_193, 2, x_192); +lean_inc(x_126); +x_194 = l_Lean_Syntax_node1(x_126, x_135, x_193); +lean_inc(x_126); +x_195 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_195, 0, x_126); +lean_ctor_set(x_195, 1, x_171); +x_196 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__28; +x_197 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__30; +x_198 = l_Lean_addMacroScope(x_31, x_197, x_32); +x_199 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__34; +lean_inc(x_126); +x_200 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_200, 0, x_126); +lean_ctor_set(x_200, 1, x_196); +lean_ctor_set(x_200, 2, x_198); +lean_ctor_set(x_200, 3, x_199); +x_201 = l_Lean_Syntax_node4(x_126, x_160, x_188, x_194, x_195, x_200); +x_202 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___closed__35; +x_203 = lean_array_push(x_202, x_173); +x_204 = lean_array_push(x_203, x_187); +x_205 = lean_array_push(x_204, x_201); +x_206 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_206, 0, x_205); +return x_206; +} +} +else +{ +uint8_t x_207; +lean_dec(x_43); +lean_dec(x_41); +lean_dec(x_39); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +lean_dec(x_3); +x_207 = !lean_is_exclusive(x_44); +if (x_207 == 0) +{ +return x_44; +} +else +{ +lean_object* x_208; lean_object* x_209; +x_208 = lean_ctor_get(x_44, 0); +lean_inc(x_208); +lean_dec(x_44); +x_209 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_209, 0, x_208); +return x_209; +} +} +} +else +{ +uint8_t x_210; +lean_dec(x_41); +lean_dec(x_39); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_210 = !lean_is_exclusive(x_42); +if (x_210 == 0) +{ +return x_42; +} +else +{ +lean_object* x_211; lean_object* x_212; +x_211 = lean_ctor_get(x_42, 0); +lean_inc(x_211); +lean_dec(x_42); +x_212 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_212, 0, x_211); +return x_212; +} +} +} +else +{ +uint8_t x_213; +lean_dec(x_39); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_213 = !lean_is_exclusive(x_40); +if (x_213 == 0) +{ +return x_40; +} +else +{ +lean_object* x_214; lean_object* x_215; +x_214 = lean_ctor_get(x_40, 0); +lean_inc(x_214); +lean_dec(x_40); +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_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_216 = !lean_is_exclusive(x_38); +if (x_216 == 0) +{ +return x_38; +} +else +{ +lean_object* x_217; lean_object* x_218; +x_217 = lean_ctor_get(x_38, 0); +lean_inc(x_217); +lean_dec(x_38); +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_dec(x_27); +lean_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_20); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_219 = !lean_is_exclusive(x_28); +if (x_219 == 0) +{ +return x_28; +} +else +{ +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_28, 0); +lean_inc(x_220); +lean_dec(x_28); +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_dec(x_25); +lean_dec(x_23); +lean_dec_ref(x_20); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_222 = !lean_is_exclusive(x_26); +if (x_222 == 0) +{ +return x_26; +} +else +{ +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_26, 0); +lean_inc(x_223); +lean_dec(x_26); +x_224 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_224, 0, x_223); +return x_224; +} +} +} +else +{ +uint8_t x_225; +lean_dec(x_23); +lean_dec_ref(x_20); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_225 = !lean_is_exclusive(x_24); +if (x_225 == 0) +{ +return x_24; +} +else +{ +lean_object* x_226; lean_object* x_227; +x_226 = lean_ctor_get(x_24, 0); +lean_inc(x_226); +lean_dec(x_24); +x_227 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_227, 0, x_226); +return x_227; +} +} +} +else +{ +uint8_t x_228; +lean_dec_ref(x_20); +lean_dec_ref(x_19); +lean_dec_ref(x_18); +lean_dec_ref(x_17); +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_10); +lean_dec(x_3); +lean_dec_ref(x_2); +x_228 = !lean_is_exclusive(x_22); +if (x_228 == 0) +{ +return x_22; +} +else +{ +lean_object* x_229; lean_object* x_230; +x_229 = lean_ctor_get(x_22, 0); +lean_inc(x_229); +lean_dec(x_22); +x_230 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_230, 0, x_229); +return x_230; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___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: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_13; +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); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_5); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_4, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +lean_dec_ref(x_4); +lean_inc_ref(x_10); +lean_inc_ref(x_6); +lean_inc(x_14); +x_16 = l_Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__0(x_14, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec_ref(x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc_ref(x_18); +x_19 = lean_ctor_get(x_17, 4); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 2); +lean_inc_ref(x_20); +lean_dec_ref(x_18); +x_21 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__0___boxed), 8, 0); +x_22 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__1___boxed), 7, 0); +lean_inc_ref(x_3); +lean_inc_ref(x_2); +lean_inc_ref(x_1); +x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___boxed), 16, 7); +lean_closure_set(x_23, 0, x_1); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_14); +lean_closure_set(x_23, 3, x_19); +lean_closure_set(x_23, 4, x_21); +lean_closure_set(x_23, 5, x_2); +lean_closure_set(x_23, 6, x_3); +x_24 = 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_25 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__17___redArg(x_20, x_23, x_24, x_24, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec_ref(x_25); +x_27 = lean_array_size(x_26); +x_28 = 0; +x_29 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__18(x_27, x_28, x_26); +x_30 = l_Array_append___redArg(x_5, x_29); +lean_dec_ref(x_29); +x_4 = x_15; +x_5 = x_30; +goto _start; +} +else +{ +lean_dec(x_15); +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); +lean_dec_ref(x_5); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +return x_25; +} +} +else +{ +uint8_t x_32; +lean_dec(x_15); +lean_dec(x_14); +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); +lean_dec_ref(x_5); +lean_dec_ref(x_3); +lean_dec_ref(x_2); +lean_dec_ref(x_1); +x_32 = !lean_is_exclusive(x_16); +if (x_32 == 0) +{ +return x_16; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_16, 0); +lean_inc(x_33); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_15; +x_15 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg(x_1, x_2, x_3, x_5, x_6, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___redArg___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_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_232; +lean_inc(x_15); +lean_inc_ref(x_14); +x_232 = l_Lean_Core_betaReduce(x_9, x_14, x_15); +if (lean_obj_tag(x_232) == 0) +{ +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_265; uint8_t x_266; +x_233 = lean_ctor_get(x_232, 0); +lean_inc(x_233); +lean_dec_ref(x_232); +x_253 = lean_ctor_get(x_1, 1); +x_254 = lean_ctor_get(x_1, 2); +x_265 = lean_unsigned_to_nat(0u); +x_266 = lean_nat_dec_lt(x_265, x_254); +if (x_266 == 0) +{ +lean_inc_ref(x_7); +x_255 = x_7; +x_256 = lean_box(0); +goto block_264; +} +else +{ +lean_object* x_267; lean_object* x_268; +lean_inc_ref(x_7); +x_267 = l_Std_Rxo_Iterator_instIteratorLoop_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__16___redArg(x_254, x_7, x_265, x_14); +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +lean_dec_ref(x_267); +x_255 = x_268; +x_256 = lean_box(0); +goto block_264; +} +block_252: +{ +lean_object* x_239; uint8_t x_240; +x_239 = lean_unsigned_to_nat(0u); +x_240 = lean_nat_dec_lt(x_239, x_4); +if (x_240 == 0) +{ +lean_dec(x_233); +lean_dec_ref(x_6); x_17 = x_235; x_18 = x_236; x_19 = x_237; @@ -4984,145 +6064,6 @@ return x_230; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___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: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_13; -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); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_5); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_4, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 1); -lean_inc(x_15); -lean_dec_ref(x_4); -lean_inc_ref(x_10); -lean_inc_ref(x_6); -lean_inc(x_14); -x_16 = l_Lean_getConstInfoCtor___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__0(x_14, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec_ref(x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc_ref(x_18); -x_19 = lean_ctor_get(x_17, 4); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 2); -lean_inc_ref(x_20); -lean_dec_ref(x_18); -x_21 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__0___boxed), 8, 0); -x_22 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__1___boxed), 7, 0); -lean_inc_ref(x_3); -lean_inc_ref(x_2); -lean_inc_ref(x_1); -x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___boxed), 16, 7); -lean_closure_set(x_23, 0, x_1); -lean_closure_set(x_23, 1, x_22); -lean_closure_set(x_23, 2, x_14); -lean_closure_set(x_23, 3, x_19); -lean_closure_set(x_23, 4, x_21); -lean_closure_set(x_23, 5, x_2); -lean_closure_set(x_23, 6, x_3); -x_24 = 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_25 = l_Lean_Meta_forallTelescopeReducing___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__17___redArg(x_20, x_23, x_24, x_24, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec_ref(x_25); -x_27 = lean_array_size(x_26); -x_28 = 0; -x_29 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__18(x_27, x_28, x_26); -x_30 = l_Array_append___redArg(x_5, x_29); -lean_dec_ref(x_29); -x_4 = x_15; -x_5 = x_30; -goto _start; -} -else -{ -lean_dec(x_15); -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); -lean_dec_ref(x_5); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -return x_25; -} -} -else -{ -uint8_t x_32; -lean_dec(x_15); -lean_dec(x_14); -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); -lean_dec_ref(x_5); -lean_dec_ref(x_3); -lean_dec_ref(x_2); -lean_dec_ref(x_1); -x_32 = !lean_is_exclusive(x_16); -if (x_32 == 0) -{ -return x_16; -} -else -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_16, 0); -lean_inc(x_33); -lean_dec(x_16); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -return x_34; -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_15; -x_15 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg(x_1, x_2, x_3, x_5, x_6, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; -} -} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___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) { _start: { @@ -5173,7 +6114,7 @@ x_23 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_ lean_inc_ref(x_3); lean_inc_ref(x_2); lean_inc_ref(x_1); -x_24 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19_spec__19___redArg___lam__2___boxed), 16, 7); +x_24 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___redArg___lam__2___boxed), 16, 7); lean_closure_set(x_24, 0, x_1); lean_closure_set(x_24, 1, x_23); lean_closure_set(x_24, 2, x_15); @@ -5764,6 +6705,17 @@ lean_dec(x_4); return x_15; } } +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___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_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___redArg___lam__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec_ref(x_8); +lean_dec(x_4); +lean_dec_ref(x_1); +return x_17; +} +} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___00__private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkMatchOld_mkAlts_spec__19___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) { _start: { diff --git a/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c index 462f7cb8176e..dad6a090f222 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c @@ -8765,8 +8765,8 @@ lean_inc(x_69); x_74 = l_Lean_Syntax_node1(x_69, x_72, x_73); x_75 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAppNTerm_spec__0___redArg___closed__16; x_76 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; -x_77 = l_Array_append___redArg(x_76, x_37); -lean_dec_ref(x_37); +x_77 = l_Array_append___redArg(x_76, x_36); +lean_dec_ref(x_36); lean_inc(x_69); x_78 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_78, 0, x_69); @@ -8919,7 +8919,7 @@ else uint8_t x_140; lean_dec(x_62); lean_dec(x_59); -lean_dec_ref(x_37); +lean_dec_ref(x_36); lean_dec_ref(x_16); lean_dec(x_33); lean_dec(x_18); @@ -8952,7 +8952,7 @@ else { uint8_t x_143; lean_dec(x_59); -lean_dec_ref(x_37); +lean_dec_ref(x_36); lean_dec_ref(x_16); lean_dec(x_33); lean_dec_ref(x_29); @@ -8985,7 +8985,7 @@ return x_145; else { uint8_t x_146; -lean_dec_ref(x_37); +lean_dec_ref(x_36); lean_dec_ref(x_16); lean_dec(x_33); lean_dec_ref(x_29); @@ -9033,8 +9033,8 @@ x_158 = lean_nat_dec_le(x_32, x_157); if (x_158 == 0) { lean_dec(x_32); -x_36 = lean_box(0); -x_37 = x_155; +x_36 = x_155; +x_37 = lean_box(0); x_38 = x_152; x_39 = x_156; x_40 = x_157; @@ -9043,8 +9043,8 @@ goto block_149; else { lean_dec(x_157); -x_36 = lean_box(0); -x_37 = x_155; +x_36 = x_155; +x_37 = lean_box(0); x_38 = x_152; x_39 = x_156; x_40 = x_32; @@ -9245,8 +9245,8 @@ lean_inc(x_210); x_215 = l_Lean_Syntax_node1(x_210, x_213, x_214); x_216 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAppNTerm_spec__0___redArg___closed__16; x_217 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; -x_218 = l_Array_append___redArg(x_217, x_178); -lean_dec_ref(x_178); +x_218 = l_Array_append___redArg(x_217, x_177); +lean_dec_ref(x_177); lean_inc(x_210); x_219 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_219, 0, x_210); @@ -9399,7 +9399,7 @@ else lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_dec(x_203); lean_dec(x_200); -lean_dec_ref(x_178); +lean_dec_ref(x_177); lean_dec_ref(x_176); lean_dec(x_173); lean_dec(x_18); @@ -9433,7 +9433,7 @@ else { lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_dec(x_200); -lean_dec_ref(x_178); +lean_dec_ref(x_177); lean_dec_ref(x_176); lean_dec(x_173); lean_dec_ref(x_169); @@ -9467,7 +9467,7 @@ return x_286; else { lean_object* x_287; lean_object* x_288; lean_object* x_289; -lean_dec_ref(x_178); +lean_dec_ref(x_177); lean_dec_ref(x_176); lean_dec(x_173); lean_dec_ref(x_169); @@ -9516,8 +9516,8 @@ x_299 = lean_nat_dec_le(x_172, x_298); if (x_299 == 0) { lean_dec(x_172); -x_177 = lean_box(0); -x_178 = x_296; +x_177 = x_296; +x_178 = lean_box(0); x_179 = x_293; x_180 = x_297; x_181 = x_298; @@ -9526,8 +9526,8 @@ goto block_290; else { lean_dec(x_298); -x_177 = lean_box(0); -x_178 = x_296; +x_177 = x_296; +x_178 = lean_box(0); x_179 = x_293; x_180 = x_297; x_181 = x_172; @@ -11453,7 +11453,7 @@ return x_216; block_153: { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_object* x_48; -x_42 = lean_array_pop(x_38); +x_42 = lean_array_pop(x_35); x_43 = l_Array_append___redArg(x_42, x_31); lean_dec(x_31); x_44 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__1; @@ -11465,22 +11465,22 @@ lean_dec_ref(x_47); if (x_12 == 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; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_49 = lean_ctor_get(x_35, 5); +x_49 = lean_ctor_get(x_37, 5); lean_inc(x_49); -x_50 = lean_ctor_get(x_35, 10); +x_50 = lean_ctor_get(x_37, 10); lean_inc(x_50); -x_51 = lean_ctor_get(x_35, 11); +x_51 = lean_ctor_get(x_37, 11); lean_inc(x_51); -lean_dec_ref(x_35); +lean_dec_ref(x_37); x_52 = l_Lean_SourceInfo_fromRef(x_49, x_12); lean_dec(x_49); x_53 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__2; x_54 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__3; -lean_inc_ref(x_39); -x_55 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_54); +lean_inc_ref(x_38); +x_55 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_54); x_56 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__4; -lean_inc_ref(x_39); -x_57 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_56); +lean_inc_ref(x_38); +x_57 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_56); x_58 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAppNTerm_spec__0___redArg___closed__16; x_59 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; lean_inc(x_52); @@ -11497,8 +11497,8 @@ lean_inc_ref_n(x_60, 7); lean_inc(x_52); x_61 = l_Lean_Syntax_node7(x_52, x_57, x_60, x_60, x_60, x_60, x_60, x_60, x_60); x_62 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__5; -lean_inc_ref(x_39); -x_63 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_62); +lean_inc_ref(x_38); +x_63 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_62); x_64 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__6; lean_inc(x_52); if (lean_is_scalar(x_32)) { @@ -11510,15 +11510,15 @@ if (lean_is_scalar(x_32)) { lean_ctor_set(x_65, 0, x_52); lean_ctor_set(x_65, 1, x_64); x_66 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__7; -lean_inc_ref(x_39); -x_67 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_66); +lean_inc_ref(x_38); +x_67 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_66); x_68 = lean_mk_syntax_ident(x_34); lean_inc_ref(x_60); lean_inc(x_52); x_69 = l_Lean_Syntax_node2(x_52, x_67, x_68, x_60); x_70 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__8; -lean_inc_ref(x_39); -x_71 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_70); +lean_inc_ref(x_38); +x_71 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_70); x_72 = l_Array_append___redArg(x_59, x_48); lean_dec_ref(x_48); lean_inc(x_52); @@ -11527,8 +11527,8 @@ lean_ctor_set(x_73, 0, x_52); lean_ctor_set(x_73, 1, x_58); lean_ctor_set(x_73, 2, x_72); x_74 = 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__9; -lean_inc_ref(x_39); -x_75 = l_Lean_Name_mkStr4(x_15, x_39, x_37, x_74); +lean_inc_ref(x_38); +x_75 = l_Lean_Name_mkStr4(x_15, x_38, x_39, x_74); x_76 = 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__11; lean_inc(x_52); x_77 = lean_alloc_ctor(2, 2, 0); @@ -11551,8 +11551,8 @@ x_84 = l_Lean_Syntax_node1(x_52, x_58, x_83); lean_inc(x_52); x_85 = l_Lean_Syntax_node2(x_52, x_71, x_73, x_84); x_86 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__17; -lean_inc_ref(x_39); -x_87 = l_Lean_Name_mkStr4(x_15, x_39, x_53, x_86); +lean_inc_ref(x_38); +x_87 = l_Lean_Name_mkStr4(x_15, x_38, x_53, x_86); x_88 = 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__27; lean_inc(x_52); x_89 = lean_alloc_ctor(2, 2, 0); @@ -11560,7 +11560,7 @@ lean_ctor_set(x_89, 0, x_52); lean_ctor_set(x_89, 1, x_88); x_90 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__18; x_91 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__19; -x_92 = l_Lean_Name_mkStr4(x_15, x_39, x_90, x_91); +x_92 = l_Lean_Name_mkStr4(x_15, x_38, x_90, x_91); lean_inc_ref_n(x_60, 2); lean_inc(x_52); x_93 = l_Lean_Syntax_node2(x_52, x_92, x_60, x_60); @@ -11581,23 +11581,23 @@ 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; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_98 = lean_ctor_get(x_35, 5); +x_98 = lean_ctor_get(x_37, 5); lean_inc(x_98); -x_99 = lean_ctor_get(x_35, 10); +x_99 = lean_ctor_get(x_37, 10); lean_inc(x_99); -x_100 = lean_ctor_get(x_35, 11); +x_100 = lean_ctor_get(x_37, 11); lean_inc(x_100); -lean_dec_ref(x_35); +lean_dec_ref(x_37); x_101 = 0; x_102 = l_Lean_SourceInfo_fromRef(x_98, x_101); lean_dec(x_98); x_103 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__2; x_104 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__3; -lean_inc_ref(x_39); -x_105 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_104); +lean_inc_ref(x_38); +x_105 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_104); x_106 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__4; -lean_inc_ref(x_39); -x_107 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_106); +lean_inc_ref(x_38); +x_107 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_106); x_108 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAppNTerm_spec__0___redArg___closed__16; x_109 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_updateIndType___closed__8; lean_inc(x_102); @@ -11611,8 +11611,8 @@ lean_ctor_set(x_110, 0, x_102); lean_ctor_set(x_110, 1, x_108); lean_ctor_set(x_110, 2, x_109); x_111 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__20; -lean_inc_ref(x_39); -x_112 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_111); +lean_inc_ref(x_38); +x_112 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_111); lean_inc(x_102); if (lean_is_scalar(x_32)) { x_113 = lean_alloc_ctor(2, 2, 0); @@ -11630,23 +11630,23 @@ lean_inc_ref_n(x_110, 6); lean_inc(x_102); x_116 = l_Lean_Syntax_node7(x_102, x_107, x_110, x_110, x_110, x_110, x_110, x_110, x_115); x_117 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__5; -lean_inc_ref(x_39); -x_118 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_117); +lean_inc_ref(x_38); +x_118 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_117); x_119 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__6; lean_inc(x_102); x_120 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_120, 0, x_102); lean_ctor_set(x_120, 1, x_119); x_121 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__7; -lean_inc_ref(x_39); -x_122 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_121); +lean_inc_ref(x_38); +x_122 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_121); x_123 = lean_mk_syntax_ident(x_34); lean_inc_ref(x_110); lean_inc(x_102); x_124 = l_Lean_Syntax_node2(x_102, x_122, x_123, x_110); x_125 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__8; -lean_inc_ref(x_39); -x_126 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_125); +lean_inc_ref(x_38); +x_126 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_125); x_127 = l_Array_append___redArg(x_109, x_48); lean_dec_ref(x_48); lean_inc(x_102); @@ -11655,8 +11655,8 @@ lean_ctor_set(x_128, 0, x_102); lean_ctor_set(x_128, 1, x_108); lean_ctor_set(x_128, 2, x_127); x_129 = 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__9; -lean_inc_ref(x_39); -x_130 = l_Lean_Name_mkStr4(x_15, x_39, x_37, x_129); +lean_inc_ref(x_38); +x_130 = l_Lean_Name_mkStr4(x_15, x_38, x_39, x_129); x_131 = 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__11; lean_inc(x_102); x_132 = lean_alloc_ctor(2, 2, 0); @@ -11679,8 +11679,8 @@ x_139 = l_Lean_Syntax_node1(x_102, x_108, x_138); lean_inc(x_102); x_140 = l_Lean_Syntax_node2(x_102, x_126, x_128, x_139); x_141 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__17; -lean_inc_ref(x_39); -x_142 = l_Lean_Name_mkStr4(x_15, x_39, x_103, x_141); +lean_inc_ref(x_38); +x_142 = l_Lean_Name_mkStr4(x_15, x_38, x_103, x_141); 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_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls_spec__1_spec__1___closed__27; lean_inc(x_102); x_144 = lean_alloc_ctor(2, 2, 0); @@ -11688,7 +11688,7 @@ lean_ctor_set(x_144, 0, x_102); lean_ctor_set(x_144, 1, x_143); x_145 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__18; x_146 = l___private_Lean_Elab_Deriving_ToExpr_0__Lean_Elab_Deriving_ToExpr_mkAuxFunction___closed__19; -x_147 = l_Lean_Name_mkStr4(x_15, x_39, x_145, x_146); +x_147 = l_Lean_Name_mkStr4(x_15, x_38, x_145, x_146); lean_inc_ref_n(x_110, 2); lean_inc(x_102); x_148 = l_Lean_Syntax_node2(x_102, x_147, x_110, x_110); @@ -11763,11 +11763,11 @@ lean_dec(x_168); lean_dec(x_166); lean_dec_ref(x_165); lean_dec(x_164); -x_154 = x_167; +x_154 = x_170; x_155 = x_162; -x_156 = x_172; -x_157 = x_170; -x_158 = x_171; +x_156 = x_167; +x_157 = x_171; +x_158 = x_172; x_159 = x_178; goto block_161; } @@ -11789,11 +11789,11 @@ lean_dec(x_168); lean_dec(x_166); lean_dec_ref(x_165); lean_dec(x_164); -x_154 = x_167; +x_154 = x_170; x_155 = x_162; -x_156 = x_172; -x_157 = x_170; -x_158 = x_171; +x_156 = x_167; +x_157 = x_171; +x_158 = x_172; x_159 = x_182; goto block_161; } @@ -11817,11 +11817,11 @@ lean_dec(x_168); lean_dec(x_166); lean_dec_ref(x_165); lean_dec(x_164); -x_154 = x_167; +x_154 = x_170; x_155 = x_162; -x_156 = x_172; -x_157 = x_170; -x_158 = x_171; +x_156 = x_167; +x_157 = x_171; +x_158 = x_172; x_159 = x_187; goto block_161; } @@ -11845,11 +11845,11 @@ lean_dec(x_168); lean_dec(x_166); lean_dec_ref(x_165); lean_dec(x_164); -x_154 = x_167; +x_154 = x_170; x_155 = x_162; -x_156 = x_172; -x_157 = x_170; -x_158 = x_171; +x_156 = x_167; +x_157 = x_171; +x_158 = x_172; x_159 = x_193; goto block_161; } @@ -11901,11 +11901,11 @@ x_211 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_211, 0, x_200); lean_ctor_set(x_211, 1, x_210); x_212 = l_Lean_Syntax_node5(x_200, x_175, x_202, x_204, x_207, x_209, x_211); -x_35 = x_167; +x_35 = x_170; x_36 = x_162; -x_37 = x_172; -x_38 = x_170; -x_39 = x_171; +x_37 = x_167; +x_38 = x_171; +x_39 = x_172; x_40 = x_212; x_41 = lean_box(0); goto block_153; @@ -11913,11 +11913,11 @@ goto block_153; else { lean_dec(x_179); -x_154 = x_167; +x_154 = x_170; x_155 = x_162; -x_156 = x_172; -x_157 = x_170; -x_158 = x_171; +x_156 = x_167; +x_157 = x_171; +x_158 = x_172; x_159 = x_195; goto block_161; } diff --git a/stage0/stdlib/Lean/Elab/Do.c b/stage0/stdlib/Lean/Elab/Do.c index 760f7305a33f..d39b5757f130 100644 --- a/stage0/stdlib/Lean/Elab/Do.c +++ b/stage0/stdlib/Lean/Elab/Do.c @@ -40547,7 +40547,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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; uint8_t 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +uint8_t x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; x_51 = lean_unsigned_to_nat(1u); x_52 = l_Lean_Syntax_getArg(x_1, x_51); x_53 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems(x_52); @@ -40581,7 +40581,7 @@ lean_inc_ref(x_3); x_61 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_Do_ToCodeBlock_doTryToCode_spec__0(x_51, x_59, x_60, x_58, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_61) == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; uint8_t 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_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_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; uint8_t x_176; 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; uint8_t x_193; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; uint8_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; 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; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; 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_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; uint8_t x_176; 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; uint8_t x_193; x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); lean_dec_ref(x_61); @@ -40658,22 +40658,22 @@ goto block_192; block_103: { 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; uint8_t x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_79 = lean_ctor_get(x_77, 0); -x_80 = lean_ctor_get(x_77, 1); -x_81 = lean_ctor_get(x_77, 2); -x_82 = lean_ctor_get(x_77, 3); -x_83 = lean_ctor_get(x_77, 4); -x_84 = lean_ctor_get(x_77, 5); -x_85 = lean_ctor_get(x_77, 6); -x_86 = lean_ctor_get(x_77, 7); -x_87 = lean_ctor_get(x_77, 8); -x_88 = lean_ctor_get(x_77, 9); -x_89 = lean_ctor_get(x_77, 10); -x_90 = lean_ctor_get(x_77, 11); -x_91 = lean_ctor_get_uint8(x_77, sizeof(void*)*14); -x_92 = lean_ctor_get(x_77, 12); -x_93 = lean_ctor_get_uint8(x_77, sizeof(void*)*14 + 1); -x_94 = lean_ctor_get(x_77, 13); +x_79 = lean_ctor_get(x_73, 0); +x_80 = lean_ctor_get(x_73, 1); +x_81 = lean_ctor_get(x_73, 2); +x_82 = lean_ctor_get(x_73, 3); +x_83 = lean_ctor_get(x_73, 4); +x_84 = lean_ctor_get(x_73, 5); +x_85 = lean_ctor_get(x_73, 6); +x_86 = lean_ctor_get(x_73, 7); +x_87 = lean_ctor_get(x_73, 8); +x_88 = lean_ctor_get(x_73, 9); +x_89 = lean_ctor_get(x_73, 10); +x_90 = lean_ctor_get(x_73, 11); +x_91 = lean_ctor_get_uint8(x_73, sizeof(void*)*14); +x_92 = lean_ctor_get(x_73, 12); +x_93 = lean_ctor_get_uint8(x_73, sizeof(void*)*14 + 1); +x_94 = lean_ctor_get(x_73, 13); x_95 = l_Lean_replaceRef(x_64, x_84); lean_dec(x_64); lean_inc_ref(x_94); @@ -40710,18 +40710,18 @@ if (x_78 == 0) { lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_dec_ref(x_77); -lean_dec_ref(x_75); -lean_dec_ref(x_74); +lean_dec(x_76); lean_dec_ref(x_73); -lean_dec_ref(x_72); -lean_dec(x_65); +lean_dec_ref(x_70); +lean_dec_ref(x_69); +lean_dec_ref(x_68); lean_dec(x_2); x_97 = l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___closed__1; -x_98 = l_Lean_throwError___at___00Lean_Elab_Term_Do_ToCodeBlock_checkReassignable_spec__0___redArg(x_97, x_67, x_70, x_96, x_76); -lean_dec(x_76); +x_98 = l_Lean_throwError___at___00Lean_Elab_Term_Do_ToCodeBlock_checkReassignable_spec__0___redArg(x_97, x_65, x_75, x_96, x_66); +lean_dec(x_66); lean_dec_ref(x_96); -lean_dec(x_70); -lean_dec_ref(x_67); +lean_dec(x_75); +lean_dec_ref(x_65); x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); lean_dec_ref(x_98); @@ -40733,57 +40733,57 @@ else { lean_object* x_101; lean_object* x_102; x_101 = lean_box(0); +lean_inc(x_66); +lean_inc(x_75); +lean_inc_ref(x_65); lean_inc(x_76); -lean_inc(x_70); -lean_inc_ref(x_67); -lean_inc(x_65); -lean_inc_ref(x_72); -lean_inc_ref(x_74); -x_102 = lean_apply_9(x_73, x_101, x_74, x_72, x_65, x_67, x_70, x_96, x_76, lean_box(0)); -x_35 = x_65; -x_36 = x_66; -x_37 = x_67; -x_38 = x_74; -x_39 = x_75; -x_40 = x_69; +lean_inc_ref(x_77); +lean_inc_ref(x_68); +x_102 = lean_apply_9(x_69, x_101, x_68, x_77, x_76, x_65, x_75, x_96, x_66, lean_box(0)); +x_35 = x_73; +x_36 = x_65; +x_37 = x_66; +x_38 = x_68; +x_39 = x_70; +x_40 = x_74; x_41 = x_71; -x_42 = x_70; +x_42 = x_75; x_43 = x_76; -x_44 = x_77; -x_45 = x_72; +x_44 = x_72; +x_45 = x_77; x_46 = x_102; goto block_50; } } block_126: { -if (lean_obj_tag(x_112) == 0) +if (lean_obj_tag(x_108) == 0) { lean_dec(x_105); lean_dec(x_104); lean_dec(x_64); lean_dec(x_1); -x_11 = x_107; +x_11 = x_112; x_12 = x_111; -x_13 = x_110; -x_14 = x_113; +x_13 = x_113; +x_14 = x_116; x_15 = x_118; -x_16 = x_109; +x_16 = x_110; x_17 = x_117; -x_18 = x_106; -x_19 = x_108; +x_18 = x_115; +x_19 = x_107; x_20 = x_114; -x_21 = x_116; -x_22 = x_115; +x_21 = x_106; +x_22 = x_109; x_23 = lean_box(0); goto block_34; } else { lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_120 = lean_ctor_get(x_112, 0); +x_120 = lean_ctor_get(x_108, 0); lean_inc(x_120); -lean_dec_ref(x_112); +lean_dec_ref(x_108); x_121 = lean_ctor_get(x_120, 0); lean_inc_ref(x_121); x_122 = lean_ctor_get(x_120, 1); @@ -40791,8 +40791,8 @@ lean_inc(x_122); lean_dec(x_120); x_123 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lam__0___boxed), 14, 5); lean_closure_set(x_123, 0, x_121); -lean_closure_set(x_123, 1, x_105); -lean_closure_set(x_123, 2, x_104); +lean_closure_set(x_123, 1, x_104); +lean_closure_set(x_123, 2, x_105); lean_closure_set(x_123, 3, x_118); lean_closure_set(x_123, 4, x_1); if (lean_obj_tag(x_122) == 0) @@ -40800,19 +40800,19 @@ if (lean_obj_tag(x_122) == 0) uint8_t x_124; lean_dec_ref(x_122); x_124 = 0; -x_65 = x_106; -x_66 = x_107; -x_67 = x_108; -x_68 = lean_box(0); -x_69 = x_111; -x_70 = x_114; +x_65 = x_107; +x_66 = x_109; +x_67 = lean_box(0); +x_68 = x_110; +x_69 = x_123; +x_70 = x_111; x_71 = x_113; -x_72 = x_117; -x_73 = x_123; -x_74 = x_109; -x_75 = x_110; +x_72 = x_116; +x_73 = x_106; +x_74 = x_112; +x_75 = x_114; x_76 = x_115; -x_77 = x_116; +x_77 = x_117; x_78 = x_124; goto block_103; } @@ -40820,19 +40820,19 @@ else { uint8_t x_125; x_125 = 1; -x_65 = x_106; -x_66 = x_107; -x_67 = x_108; -x_68 = lean_box(0); -x_69 = x_111; -x_70 = x_114; +x_65 = x_107; +x_66 = x_109; +x_67 = lean_box(0); +x_68 = x_110; +x_69 = x_123; +x_70 = x_111; x_71 = x_113; -x_72 = x_117; -x_73 = x_123; -x_74 = x_109; -x_75 = x_110; +x_72 = x_116; +x_73 = x_106; +x_74 = x_112; +x_75 = x_114; x_76 = x_115; -x_77 = x_116; +x_77 = x_117; x_78 = x_125; goto block_103; } @@ -40909,19 +40909,19 @@ if (x_157 == 0) lean_dec(x_156); lean_dec(x_136); lean_dec(x_62); -x_104 = x_141; -x_105 = x_140; -x_106 = x_130; -x_107 = x_148; -x_108 = x_131; -x_109 = x_128; -x_110 = x_142; -x_111 = x_144; -x_112 = x_127; -x_113 = x_146; +x_104 = x_140; +x_105 = x_141; +x_106 = x_133; +x_107 = x_131; +x_108 = x_127; +x_109 = x_134; +x_110 = x_128; +x_111 = x_142; +x_112 = x_146; +x_113 = x_144; x_114 = x_132; -x_115 = x_134; -x_116 = x_133; +x_115 = x_130; +x_116 = x_148; x_117 = x_129; x_118 = x_154; x_119 = lean_box(0); @@ -40936,19 +40936,19 @@ if (x_158 == 0) lean_dec(x_156); lean_dec(x_136); lean_dec(x_62); -x_104 = x_141; -x_105 = x_140; -x_106 = x_130; -x_107 = x_148; -x_108 = x_131; -x_109 = x_128; -x_110 = x_142; -x_111 = x_144; -x_112 = x_127; -x_113 = x_146; +x_104 = x_140; +x_105 = x_141; +x_106 = x_133; +x_107 = x_131; +x_108 = x_127; +x_109 = x_134; +x_110 = x_128; +x_111 = x_142; +x_112 = x_146; +x_113 = x_144; x_114 = x_132; -x_115 = x_134; -x_116 = x_133; +x_115 = x_130; +x_116 = x_148; x_117 = x_129; x_118 = x_154; x_119 = lean_box(0); @@ -40976,19 +40976,19 @@ lean_object* x_161; x_161 = lean_ctor_get(x_160, 0); lean_inc(x_161); lean_dec_ref(x_160); -x_104 = x_141; -x_105 = x_140; -x_106 = x_130; -x_107 = x_148; -x_108 = x_131; -x_109 = x_128; -x_110 = x_142; -x_111 = x_144; -x_112 = x_127; -x_113 = x_146; +x_104 = x_140; +x_105 = x_141; +x_106 = x_133; +x_107 = x_131; +x_108 = x_127; +x_109 = x_134; +x_110 = x_128; +x_111 = x_142; +x_112 = x_146; +x_113 = x_144; x_114 = x_132; -x_115 = x_134; -x_116 = x_133; +x_115 = x_130; +x_116 = x_148; x_117 = x_129; x_118 = x_161; x_119 = lean_box(0); @@ -41071,14 +41071,14 @@ return x_137; { if (x_176 == 0) { -x_127 = x_170; -x_128 = x_168; +x_127 = x_173; +x_128 = x_170; x_129 = x_175; -x_130 = x_167; -x_131 = x_174; -x_132 = x_171; -x_133 = x_169; -x_134 = x_173; +x_130 = x_169; +x_131 = x_168; +x_132 = x_174; +x_133 = x_172; +x_134 = x_167; x_135 = lean_box(0); goto block_166; } @@ -41086,20 +41086,20 @@ else { lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_dec_ref(x_175); -lean_dec(x_170); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_173); +lean_dec_ref(x_170); +lean_dec(x_169); lean_dec(x_64); lean_dec(x_62); lean_dec(x_55); lean_dec(x_2); lean_dec(x_1); x_177 = l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___closed__6; -x_178 = l_Lean_throwError___at___00Lean_Elab_Term_Do_ToCodeBlock_checkReassignable_spec__0___redArg(x_177, x_174, x_171, x_169, x_173); -lean_dec(x_173); -lean_dec_ref(x_169); -lean_dec(x_171); -lean_dec_ref(x_174); +x_178 = l_Lean_throwError___at___00Lean_Elab_Term_Do_ToCodeBlock_checkReassignable_spec__0___redArg(x_177, x_168, x_174, x_172, x_167); +lean_dec(x_167); +lean_dec_ref(x_172); +lean_dec(x_174); +lean_dec_ref(x_168); x_179 = lean_ctor_get(x_178, 0); lean_inc(x_179); lean_dec_ref(x_178); @@ -41114,14 +41114,14 @@ uint8_t x_191; x_191 = l_Array_isEmpty___redArg(x_62); if (x_191 == 0) { -x_167 = x_185; -x_168 = x_183; -x_169 = x_188; -x_170 = x_182; -x_171 = x_187; -x_172 = lean_box(0); -x_173 = x_189; -x_174 = x_186; +x_167 = x_189; +x_168 = x_186; +x_169 = x_185; +x_170 = x_183; +x_171 = lean_box(0); +x_172 = x_188; +x_173 = x_182; +x_174 = x_187; x_175 = x_184; x_176 = x_191; goto block_181; @@ -41130,14 +41130,14 @@ else { if (lean_obj_tag(x_182) == 0) { -x_167 = x_185; -x_168 = x_183; -x_169 = x_188; -x_170 = x_182; -x_171 = x_187; -x_172 = lean_box(0); -x_173 = x_189; -x_174 = x_186; +x_167 = x_189; +x_168 = x_186; +x_169 = x_185; +x_170 = x_183; +x_171 = lean_box(0); +x_172 = x_188; +x_173 = x_182; +x_174 = x_187; x_175 = x_184; x_176 = x_191; goto block_181; @@ -41195,12 +41195,12 @@ return x_54; block_34: { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_box(x_12); -x_25 = lean_box(x_14); -x_26 = lean_box(x_11); +x_24 = lean_box(x_13); +x_25 = lean_box(x_11); +x_26 = lean_box(x_14); x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___boxed), 7, 5); lean_closure_set(x_27, 0, x_15); -lean_closure_set(x_27, 1, x_13); +lean_closure_set(x_27, 1, x_12); lean_closure_set(x_27, 2, x_24); lean_closure_set(x_27, 3, x_25); lean_closure_set(x_27, 4, x_26); @@ -41243,18 +41243,18 @@ lean_object* x_47; x_47 = lean_ctor_get(x_46, 0); lean_inc(x_47); lean_dec_ref(x_46); -x_11 = x_36; -x_12 = x_40; -x_13 = x_39; -x_14 = x_41; +x_11 = x_40; +x_12 = x_39; +x_13 = x_41; +x_14 = x_44; x_15 = x_47; x_16 = x_38; x_17 = x_45; -x_18 = x_35; -x_19 = x_37; +x_18 = x_43; +x_19 = x_36; x_20 = x_42; -x_21 = x_44; -x_22 = x_43; +x_21 = x_35; +x_22 = x_37; x_23 = lean_box(0); goto block_34; } @@ -41262,13 +41262,13 @@ else { lean_object* x_48; lean_object* x_49; lean_dec_ref(x_45); -lean_dec_ref(x_44); lean_dec(x_43); lean_dec(x_42); lean_dec_ref(x_39); lean_dec_ref(x_38); -lean_dec_ref(x_37); -lean_dec(x_35); +lean_dec(x_37); +lean_dec_ref(x_36); +lean_dec_ref(x_35); lean_dec(x_2); x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); @@ -43218,8 +43218,8 @@ x_394 = l_Lean_Elab_Term_Do_annotate___at___00Lean_Elab_Term_Do_ToCodeBlock_doTr x_395 = lean_ctor_get(x_394, 0); lean_inc(x_395); lean_dec_ref(x_394); -x_11 = x_364; -x_12 = x_359; +x_11 = x_359; +x_12 = x_364; x_13 = x_395; x_14 = x_3; x_15 = x_4; @@ -43288,8 +43288,8 @@ x_417 = l_Lean_Elab_Term_Do_annotate___at___00Lean_Elab_Term_Do_ToCodeBlock_doTr x_418 = lean_ctor_get(x_417, 0); lean_inc(x_418); lean_dec_ref(x_417); -x_11 = x_364; -x_12 = x_359; +x_11 = x_359; +x_12 = x_364; x_13 = x_418; x_14 = x_3; x_15 = x_4; @@ -44067,7 +44067,7 @@ x_22 = l_List_isEmpty___redArg(x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec_ref(x_11); +lean_dec_ref(x_12); x_23 = lean_ctor_get(x_19, 5); x_24 = lean_ctor_get(x_19, 10); x_25 = lean_ctor_get(x_19, 11); @@ -44139,7 +44139,7 @@ lean_ctor_set(x_57, 0, x_26); lean_ctor_set(x_57, 1, x_56); lean_inc_ref_n(x_37, 2); lean_inc(x_26); -x_58 = l_Lean_Syntax_node5(x_26, x_55, x_12, x_37, x_37, x_57, x_43); +x_58 = l_Lean_Syntax_node5(x_26, x_55, x_11, x_37, x_37, x_57, x_43); lean_inc(x_26); x_59 = l_Lean_Syntax_node1(x_26, x_54, x_58); lean_inc(x_26); @@ -44171,7 +44171,7 @@ lean_inc(x_16); lean_inc_ref(x_15); lean_inc_ref(x_14); lean_inc(x_68); -x_71 = lean_apply_9(x_11, x_68, x_14, x_15, x_16, x_17, x_18, x_19, x_20, lean_box(0)); +x_71 = lean_apply_9(x_12, x_68, x_14, x_15, x_16, x_17, x_18, x_19, x_20, lean_box(0)); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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; @@ -44246,7 +44246,7 @@ lean_ctor_set(x_103, 0, x_72); lean_ctor_set(x_103, 1, x_102); lean_inc_ref_n(x_83, 2); lean_inc(x_72); -x_104 = l_Lean_Syntax_node5(x_72, x_101, x_12, x_83, x_83, x_103, x_89); +x_104 = l_Lean_Syntax_node5(x_72, x_101, x_11, x_83, x_83, x_103, x_89); lean_inc(x_72); x_105 = l_Lean_Syntax_node1(x_72, x_100, x_104); lean_inc(x_72); @@ -44353,7 +44353,7 @@ lean_dec(x_16); lean_dec_ref(x_15); lean_dec_ref(x_14); lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_11); x_151 = lean_ctor_get(x_71, 0); lean_inc(x_151); lean_dec_ref(x_71); diff --git a/stage0/stdlib/Lean/Elab/ElabRules.c b/stage0/stdlib/Lean/Elab/ElabRules.c index 6f136076b89d..9fcbb4bdfeb0 100644 --- a/stage0/stdlib/Lean/Elab/ElabRules.c +++ b/stage0/stdlib/Lean/Elab/ElabRules.c @@ -1597,9 +1597,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_62; -x_26 = lean_box(0); -x_27 = x_63; +x_25 = x_63; +x_26 = x_62; +x_27 = lean_box(0); goto block_34; } else @@ -1615,9 +1615,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_62; -x_26 = lean_box(0); -x_27 = x_63; +x_25 = x_63; +x_26 = x_62; +x_27 = lean_box(0); goto block_34; } else @@ -1781,7 +1781,7 @@ x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__priv 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_throwErrorAt___at___00Lean_Elab_Command_elabElabRulesAux_spec__3___redArg(x_10, x_32, x_25, x_27); +x_33 = l_Lean_throwErrorAt___at___00Lean_Elab_Command_elabElabRulesAux_spec__3___redArg(x_10, x_32, x_26, x_25); lean_dec(x_10); x_20 = x_33; goto block_24; @@ -1904,27 +1904,27 @@ goto block_99; { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; x_48 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__9; -lean_inc(x_45); +lean_inc(x_46); x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_45); +lean_ctor_set(x_49, 0, x_46); lean_ctor_set(x_49, 1, x_48); x_50 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__11; x_51 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__12; -x_52 = l_Array_append___redArg(x_51, x_46); -lean_dec_ref(x_46); -lean_inc(x_45); +x_52 = l_Array_append___redArg(x_51, x_45); +lean_dec_ref(x_45); +lean_inc(x_46); x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_45); +lean_ctor_set(x_53, 0, x_46); lean_ctor_set(x_53, 1, x_50); lean_ctor_set(x_53, 2, x_52); -lean_inc(x_45); -x_54 = l_Lean_Syntax_node1(x_45, x_50, x_53); +lean_inc(x_46); +x_54 = l_Lean_Syntax_node1(x_46, x_50, x_53); x_55 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; -lean_inc(x_45); +lean_inc(x_46); x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_45); +lean_ctor_set(x_56, 0, x_46); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Syntax_node4(x_45, x_35, x_49, x_54, x_56, x_44); +x_57 = l_Lean_Syntax_node4(x_46, x_35, x_49, x_54, x_56, x_44); x_13 = x_57; x_14 = lean_box(0); goto block_19; @@ -1991,8 +1991,8 @@ lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); x_25 = lean_box(0); -x_26 = x_63; -x_27 = x_62; +x_26 = x_62; +x_27 = x_63; goto block_34; } else @@ -2009,8 +2009,8 @@ lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); x_25 = lean_box(0); -x_26 = x_63; -x_27 = x_62; +x_26 = x_62; +x_27 = x_63; goto block_34; } else @@ -2044,16 +2044,16 @@ if (lean_obj_tag(x_88) == 0) lean_object* x_92; x_92 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_63); lean_dec_ref(x_92); -x_45 = x_91; -x_46 = x_90; +x_45 = x_90; +x_46 = x_91; x_47 = lean_box(0); goto block_58; } else { lean_dec_ref(x_88); -x_45 = x_91; -x_46 = x_90; +x_45 = x_90; +x_46 = x_91; x_47 = lean_box(0); goto block_58; } @@ -2173,7 +2173,7 @@ x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__priv 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_throwErrorAt___at___00Lean_Elab_Command_elabElabRulesAux_spec__3___redArg(x_10, x_32, x_27, x_26); +x_33 = l_Lean_throwErrorAt___at___00Lean_Elab_Command_elabElabRulesAux_spec__3___redArg(x_10, x_32, x_26, x_27); lean_dec(x_10); x_20 = x_33; goto block_24; @@ -2969,7 +2969,7 @@ lean_inc(x_4); x_13 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10(x_4, x_11, x_12, x_7, x_8, x_9); 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_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_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_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; uint8_t 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_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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_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_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; uint8_t x_536; lean_object* x_537; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); if (lean_is_exclusive(x_13)) { @@ -3032,287 +3032,287 @@ goto block_640; block_151: { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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_inc_ref(x_27); -x_30 = l_Array_append___redArg(x_27, x_29); +lean_inc_ref(x_19); +x_30 = l_Array_append___redArg(x_19, x_29); lean_dec_ref(x_29); -lean_inc(x_26); -lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_17); x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_24); -lean_ctor_set(x_31, 1, x_26); +lean_ctor_set(x_31, 0, x_17); +lean_ctor_set(x_31, 1, x_23); lean_ctor_set(x_31, 2, x_30); x_32 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__5; x_33 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__6; x_34 = l_Lean_Elab_Command_elabElabRulesAux___closed__0; -lean_inc_ref(x_16); -x_35 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_34); +lean_inc_ref(x_22); +x_35 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_34); x_36 = l_Lean_Elab_Command_elabElabRulesAux___closed__1; -lean_inc(x_24); +lean_inc(x_17); x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_24); +lean_ctor_set(x_37, 0, x_17); lean_ctor_set(x_37, 1, x_36); x_38 = l_Lean_Elab_Command_elabElabRulesAux___closed__2; -x_39 = l_Lean_Syntax_SepArray_ofElems(x_38, x_20); -lean_dec_ref(x_20); -lean_inc_ref(x_27); -x_40 = l_Array_append___redArg(x_27, x_39); +x_39 = l_Lean_Syntax_SepArray_ofElems(x_38, x_24); +lean_dec_ref(x_24); +lean_inc_ref(x_19); +x_40 = l_Array_append___redArg(x_19, x_39); lean_dec_ref(x_39); -lean_inc(x_26); -lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_17); x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_24); -lean_ctor_set(x_41, 1, x_26); +lean_ctor_set(x_41, 0, x_17); +lean_ctor_set(x_41, 1, x_23); lean_ctor_set(x_41, 2, x_40); x_42 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; -lean_inc(x_24); +lean_inc(x_17); x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); +lean_ctor_set(x_43, 0, x_17); lean_ctor_set(x_43, 1, x_42); -lean_inc(x_24); -x_44 = l_Lean_Syntax_node3(x_24, x_35, x_37, x_41, x_43); -lean_inc(x_26); -lean_inc(x_24); -x_45 = l_Lean_Syntax_node1(x_24, x_26, x_44); -lean_inc(x_24); +lean_inc(x_17); +x_44 = l_Lean_Syntax_node3(x_17, x_35, x_37, x_41, x_43); +lean_inc(x_23); +lean_inc(x_17); +x_45 = l_Lean_Syntax_node1(x_17, x_23, x_44); +lean_inc(x_17); x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_24); +lean_ctor_set(x_46, 0, x_17); lean_ctor_set(x_46, 1, x_28); x_47 = l_Lean_Elab_Command_elabElabRulesAux___closed__5; x_48 = l_Lean_Elab_Command_elabElabRulesAux___closed__6; -lean_inc(x_19); lean_inc(x_18); -x_49 = l_Lean_addMacroScope(x_18, x_48, x_19); +lean_inc(x_27); +x_49 = l_Lean_addMacroScope(x_27, x_48, x_18); x_50 = lean_box(0); -lean_inc(x_24); +lean_inc(x_17); x_51 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_51, 0, x_24); +lean_ctor_set(x_51, 0, x_17); lean_ctor_set(x_51, 1, x_47); lean_ctor_set(x_51, 2, x_49); lean_ctor_set(x_51, 3, x_50); x_52 = lean_mk_syntax_ident(x_4); -lean_inc(x_26); -lean_inc(x_24); -x_53 = l_Lean_Syntax_node2(x_24, x_26, x_51, x_52); +lean_inc(x_23); +lean_inc(x_17); +x_53 = l_Lean_Syntax_node2(x_17, x_23, x_51, x_52); x_54 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_24); +lean_inc(x_17); x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_24); +lean_ctor_set(x_55, 0, x_17); lean_ctor_set(x_55, 1, x_54); x_56 = l_Lean_Elab_Command_elabElabRulesAux___closed__9; x_57 = l_Lean_Elab_Command_elabElabRulesAux___closed__10; -lean_inc_ref(x_23); -lean_inc_ref(x_16); -x_58 = l_Lean_Name_mkStr4(x_16, x_23, x_33, x_57); -lean_inc(x_19); -lean_inc(x_58); +lean_inc_ref(x_20); +lean_inc_ref(x_22); +x_58 = l_Lean_Name_mkStr4(x_22, x_20, x_33, x_57); lean_inc(x_18); -x_59 = l_Lean_addMacroScope(x_18, x_58, x_19); +lean_inc(x_58); +lean_inc(x_27); +x_59 = l_Lean_addMacroScope(x_27, x_58, x_18); x_60 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_50); x_61 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_61, 0, x_60); lean_ctor_set(x_61, 1, x_50); -lean_inc(x_24); +lean_inc(x_17); x_62 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_62, 0, x_24); +lean_ctor_set(x_62, 0, x_17); lean_ctor_set(x_62, 1, x_56); lean_ctor_set(x_62, 2, x_59); lean_ctor_set(x_62, 3, x_61); x_63 = l_Lean_Elab_Command_elabElabRulesAux___closed__11; -lean_inc(x_24); +lean_inc(x_17); x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_24); +lean_ctor_set(x_64, 0, x_17); lean_ctor_set(x_64, 1, x_63); x_65 = l_Lean_Elab_Command_elabElabRulesAux___closed__12; -lean_inc_ref(x_16); -x_66 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_65); -lean_inc(x_24); +lean_inc_ref(x_22); +x_66 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_65); +lean_inc(x_17); x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_24); +lean_ctor_set(x_67, 0, x_17); lean_ctor_set(x_67, 1, x_65); x_68 = l_Lean_Elab_Command_elabElabRulesAux___closed__13; -lean_inc_ref(x_16); -x_69 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_68); +lean_inc_ref(x_22); +x_69 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_68); x_70 = l_Lean_Elab_Command_elabElabRulesAux___closed__15; x_71 = l_Lean_Elab_Command_elabElabRulesAux___closed__16; -lean_inc(x_19); lean_inc(x_18); -x_72 = l_Lean_addMacroScope(x_18, x_71, x_19); -lean_inc(x_24); +lean_inc(x_27); +x_72 = l_Lean_addMacroScope(x_27, x_71, x_18); +lean_inc(x_17); x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_24); +lean_ctor_set(x_73, 0, x_17); lean_ctor_set(x_73, 1, x_70); lean_ctor_set(x_73, 2, x_72); lean_ctor_set(x_73, 3, x_50); x_74 = l_Lean_Elab_Command_elabElabRulesAux___closed__18; x_75 = l_Lean_Elab_Command_elabElabRulesAux___closed__19; -lean_inc(x_19); lean_inc(x_18); -x_76 = l_Lean_addMacroScope(x_18, x_75, x_19); -lean_inc(x_24); +lean_inc(x_27); +x_76 = l_Lean_addMacroScope(x_27, x_75, x_18); +lean_inc(x_17); x_77 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_77, 0, x_24); +lean_ctor_set(x_77, 0, x_17); lean_ctor_set(x_77, 1, x_74); lean_ctor_set(x_77, 2, x_76); lean_ctor_set(x_77, 3, x_50); lean_inc_ref(x_77); lean_inc_ref(x_73); -lean_inc(x_26); -lean_inc(x_24); -x_78 = l_Lean_Syntax_node2(x_24, x_26, x_73, x_77); -lean_inc_ref(x_27); -lean_inc(x_26); -lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_17); +x_78 = l_Lean_Syntax_node2(x_17, x_23, x_73, x_77); +lean_inc_ref(x_19); +lean_inc(x_23); +lean_inc(x_17); x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_24); -lean_ctor_set(x_79, 1, x_26); -lean_ctor_set(x_79, 2, x_27); +lean_ctor_set(x_79, 0, x_17); +lean_ctor_set(x_79, 1, x_23); +lean_ctor_set(x_79, 2, x_19); x_80 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; -lean_inc(x_24); +lean_inc(x_17); x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_24); +lean_ctor_set(x_81, 0, x_17); lean_ctor_set(x_81, 1, x_80); x_82 = l_Lean_Elab_Command_elabElabRulesAux___closed__20; -lean_inc_ref(x_16); -x_83 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_82); +lean_inc_ref(x_22); +x_83 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_82); x_84 = l_Lean_Elab_Command_elabElabRulesAux___closed__22; x_85 = l_Lean_Elab_Command_elabElabRulesAux___closed__23; -lean_inc_ref(x_23); -lean_inc_ref(x_16); -x_86 = l_Lean_Name_mkStr4(x_16, x_23, x_33, x_85); -lean_inc(x_19); -lean_inc(x_86); +lean_inc_ref(x_20); +lean_inc_ref(x_22); +x_86 = l_Lean_Name_mkStr4(x_22, x_20, x_33, x_85); lean_inc(x_18); -x_87 = l_Lean_addMacroScope(x_18, x_86, x_19); +lean_inc(x_86); +lean_inc(x_27); +x_87 = l_Lean_addMacroScope(x_27, x_86, x_18); x_88 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_88, 0, x_86); lean_ctor_set(x_88, 1, x_50); x_89 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_50); -lean_inc(x_24); +lean_inc(x_17); x_90 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_90, 0, x_24); +lean_ctor_set(x_90, 0, x_17); lean_ctor_set(x_90, 1, x_84); lean_ctor_set(x_90, 2, x_87); lean_ctor_set(x_90, 3, x_89); -lean_inc(x_26); -lean_inc(x_24); -x_91 = l_Lean_Syntax_node1(x_24, x_26, x_25); +lean_inc(x_23); +lean_inc(x_17); +x_91 = l_Lean_Syntax_node1(x_17, x_23, x_21); x_92 = l_Lean_Elab_Command_elabElabRulesAux___closed__24; -lean_inc_ref(x_16); -x_93 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_92); -lean_inc(x_24); +lean_inc_ref(x_22); +x_93 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_92); +lean_inc(x_17); x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_24); +lean_ctor_set(x_94, 0, x_17); lean_ctor_set(x_94, 1, x_92); x_95 = l_Lean_Elab_Command_elabElabRulesAux___closed__25; -lean_inc_ref(x_16); -x_96 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_95); +lean_inc_ref(x_22); +x_96 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_95); lean_inc_ref(x_79); -lean_inc(x_24); -x_97 = l_Lean_Syntax_node2(x_24, x_96, x_79, x_73); -lean_inc(x_26); -lean_inc(x_24); -x_98 = l_Lean_Syntax_node1(x_24, x_26, x_97); +lean_inc(x_17); +x_97 = l_Lean_Syntax_node2(x_17, x_96, x_79, x_73); +lean_inc(x_23); +lean_inc(x_17); +x_98 = l_Lean_Syntax_node1(x_17, x_23, x_97); x_99 = l_Lean_Elab_Command_elabElabRulesAux___closed__26; -lean_inc(x_24); +lean_inc(x_17); x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_24); +lean_ctor_set(x_100, 0, x_17); lean_ctor_set(x_100, 1, x_99); x_101 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; -lean_inc_ref(x_16); -x_102 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_101); +lean_inc_ref(x_22); +x_102 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_101); x_103 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__7; -lean_inc_ref(x_16); -x_104 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_103); -x_105 = l_Array_append___redArg(x_27, x_14); +lean_inc_ref(x_22); +x_104 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_103); +x_105 = l_Array_append___redArg(x_19, x_14); lean_dec(x_14); x_106 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__9; -lean_inc(x_24); +lean_inc(x_17); x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_24); +lean_ctor_set(x_107, 0, x_17); lean_ctor_set(x_107, 1, x_106); x_108 = l_Lean_Elab_Command_elabElabRulesAux___closed__28; -lean_inc_ref(x_16); -x_109 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_108); +lean_inc_ref(x_22); +x_109 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_108); x_110 = l_Lean_Elab_Command_elabElabRulesAux___closed__29; -lean_inc(x_24); +lean_inc(x_17); x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_24); +lean_ctor_set(x_111, 0, x_17); lean_ctor_set(x_111, 1, x_110); -lean_inc(x_24); -x_112 = l_Lean_Syntax_node1(x_24, x_109, x_111); -lean_inc(x_26); -lean_inc(x_24); -x_113 = l_Lean_Syntax_node1(x_24, x_26, x_112); -lean_inc(x_26); -lean_inc(x_24); -x_114 = l_Lean_Syntax_node1(x_24, x_26, x_113); +lean_inc(x_17); +x_112 = l_Lean_Syntax_node1(x_17, x_109, x_111); +lean_inc(x_23); +lean_inc(x_17); +x_113 = l_Lean_Syntax_node1(x_17, x_23, x_112); +lean_inc(x_23); +lean_inc(x_17); +x_114 = l_Lean_Syntax_node1(x_17, x_23, x_113); x_115 = l_Lean_Elab_Command_elabElabRulesAux___closed__30; -lean_inc_ref(x_16); -x_116 = l_Lean_Name_mkStr4(x_16, x_32, x_33, x_115); +lean_inc_ref(x_22); +x_116 = l_Lean_Name_mkStr4(x_22, x_32, x_33, x_115); x_117 = l_Lean_Elab_Command_elabElabRulesAux___closed__31; -lean_inc(x_24); +lean_inc(x_17); x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_24); +lean_ctor_set(x_118, 0, x_17); lean_ctor_set(x_118, 1, x_117); x_119 = l_Lean_Elab_Command_elabElabRulesAux___closed__32; x_120 = l_Lean_Elab_Command_elabElabRulesAux___closed__33; x_121 = l_Lean_Elab_Command_elabElabRulesAux___closed__34; -x_122 = l_Lean_addMacroScope(x_18, x_121, x_19); -x_123 = l_Lean_Name_mkStr3(x_16, x_23, x_119); +x_122 = l_Lean_addMacroScope(x_27, x_121, x_18); +x_123 = l_Lean_Name_mkStr3(x_22, x_20, x_119); x_124 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_124, 0, x_123); lean_ctor_set(x_124, 1, x_50); x_125 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_125, 0, x_124); lean_ctor_set(x_125, 1, x_50); -lean_inc(x_24); +lean_inc(x_17); x_126 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_126, 0, x_24); +lean_ctor_set(x_126, 0, x_17); lean_ctor_set(x_126, 1, x_120); lean_ctor_set(x_126, 2, x_122); lean_ctor_set(x_126, 3, x_125); -lean_inc(x_24); -x_127 = l_Lean_Syntax_node2(x_24, x_116, x_118, x_126); +lean_inc(x_17); +x_127 = l_Lean_Syntax_node2(x_17, x_116, x_118, x_126); lean_inc_ref(x_81); -lean_inc(x_24); -x_128 = l_Lean_Syntax_node4(x_24, x_104, x_107, x_114, x_81, x_127); +lean_inc(x_17); +x_128 = l_Lean_Syntax_node4(x_17, x_104, x_107, x_114, x_81, x_127); x_129 = lean_array_push(x_105, x_128); -lean_inc(x_26); -lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_17); x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_24); -lean_ctor_set(x_130, 1, x_26); +lean_ctor_set(x_130, 0, x_17); +lean_ctor_set(x_130, 1, x_23); lean_ctor_set(x_130, 2, x_129); -lean_inc(x_24); -x_131 = l_Lean_Syntax_node1(x_24, x_102, x_130); +lean_inc(x_17); +x_131 = l_Lean_Syntax_node1(x_17, x_102, x_130); lean_inc_ref_n(x_79, 2); -lean_inc(x_24); -x_132 = l_Lean_Syntax_node6(x_24, x_93, x_94, x_79, x_79, x_98, x_100, x_131); +lean_inc(x_17); +x_132 = l_Lean_Syntax_node6(x_17, x_93, x_94, x_79, x_79, x_98, x_100, x_131); lean_inc_ref(x_81); lean_inc_ref(x_79); lean_inc(x_69); -lean_inc(x_24); -x_133 = l_Lean_Syntax_node4(x_24, x_69, x_91, x_79, x_81, x_132); +lean_inc(x_17); +x_133 = l_Lean_Syntax_node4(x_17, x_69, x_91, x_79, x_81, x_132); lean_inc_ref(x_67); lean_inc(x_66); -lean_inc(x_24); -x_134 = l_Lean_Syntax_node2(x_24, x_66, x_67, x_133); -lean_inc(x_24); -x_135 = l_Lean_Syntax_node2(x_24, x_26, x_77, x_134); -lean_inc(x_24); -x_136 = l_Lean_Syntax_node2(x_24, x_83, x_90, x_135); -lean_inc(x_24); -x_137 = l_Lean_Syntax_node4(x_24, x_69, x_78, x_79, x_81, x_136); -lean_inc(x_24); -x_138 = l_Lean_Syntax_node2(x_24, x_66, x_67, x_137); +lean_inc(x_17); +x_134 = l_Lean_Syntax_node2(x_17, x_66, x_67, x_133); +lean_inc(x_17); +x_135 = l_Lean_Syntax_node2(x_17, x_23, x_77, x_134); +lean_inc(x_17); +x_136 = l_Lean_Syntax_node2(x_17, x_83, x_90, x_135); +lean_inc(x_17); +x_137 = l_Lean_Syntax_node4(x_17, x_69, x_78, x_79, x_81, x_136); +lean_inc(x_17); +x_138 = l_Lean_Syntax_node2(x_17, x_66, x_67, x_137); x_139 = l_Lean_Elab_Command_elabElabRulesAux___closed__35; x_140 = lean_array_push(x_139, x_31); x_141 = lean_array_push(x_140, x_45); -x_142 = lean_array_push(x_141, x_21); +x_142 = lean_array_push(x_141, x_16); x_143 = lean_array_push(x_142, x_46); x_144 = lean_array_push(x_143, x_53); x_145 = lean_array_push(x_144, x_55); @@ -3320,8 +3320,8 @@ x_146 = lean_array_push(x_145, x_62); x_147 = lean_array_push(x_146, x_64); x_148 = lean_array_push(x_147, x_138); x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_24); -lean_ctor_set(x_149, 1, x_17); +lean_ctor_set(x_149, 0, x_17); +lean_ctor_set(x_149, 1, x_26); lean_ctor_set(x_149, 2, x_148); if (lean_is_scalar(x_15)) { x_150 = lean_alloc_ctor(0, 1, 0); @@ -3344,18 +3344,18 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_165; x_165 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_16 = x_159; -x_17 = x_162; -x_18 = x_157; -x_19 = x_153; -x_20 = x_155; -x_21 = x_154; -x_22 = lean_box(0); -x_23 = x_160; -x_24 = x_156; -x_25 = x_152; -x_26 = x_163; -x_27 = x_164; +x_16 = x_152; +x_17 = x_154; +x_18 = x_153; +x_19 = x_164; +x_20 = x_160; +x_21 = x_156; +x_22 = x_159; +x_23 = x_163; +x_24 = x_155; +x_25 = lean_box(0); +x_26 = x_162; +x_27 = x_157; x_28 = x_161; x_29 = x_165; goto block_151; @@ -3367,18 +3367,18 @@ x_166 = lean_ctor_get(x_1, 0); lean_inc(x_166); lean_dec_ref(x_1); x_167 = l_Array_mkArray1___redArg(x_166); -x_16 = x_159; -x_17 = x_162; -x_18 = x_157; -x_19 = x_153; -x_20 = x_155; -x_21 = x_154; -x_22 = lean_box(0); -x_23 = x_160; -x_24 = x_156; -x_25 = x_152; -x_26 = x_163; -x_27 = x_164; +x_16 = x_152; +x_17 = x_154; +x_18 = x_153; +x_19 = x_164; +x_20 = x_160; +x_21 = x_156; +x_22 = x_159; +x_23 = x_163; +x_24 = x_155; +x_25 = lean_box(0); +x_26 = x_162; +x_27 = x_157; x_28 = x_161; x_29 = x_167; goto block_151; @@ -3387,14 +3387,14 @@ goto block_151; block_285: { lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -lean_inc_ref(x_179); -x_182 = l_Array_append___redArg(x_179, x_181); +lean_inc_ref(x_172); +x_182 = l_Array_append___redArg(x_172, x_181); lean_dec_ref(x_181); -lean_inc(x_175); lean_inc(x_180); +lean_inc(x_179); x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_175); +lean_ctor_set(x_183, 0, x_179); +lean_ctor_set(x_183, 1, x_180); lean_ctor_set(x_183, 2, x_182); x_184 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__5; x_185 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__6; @@ -3402,101 +3402,101 @@ x_186 = l_Lean_Elab_Command_elabElabRulesAux___closed__0; lean_inc_ref(x_177); x_187 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_186); x_188 = l_Lean_Elab_Command_elabElabRulesAux___closed__1; -lean_inc(x_180); +lean_inc(x_179); x_189 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_189, 0, x_180); +lean_ctor_set(x_189, 0, x_179); lean_ctor_set(x_189, 1, x_188); x_190 = l_Lean_Elab_Command_elabElabRulesAux___closed__2; -x_191 = l_Lean_Syntax_SepArray_ofElems(x_190, x_171); -lean_dec_ref(x_171); -lean_inc_ref(x_179); -x_192 = l_Array_append___redArg(x_179, x_191); +x_191 = l_Lean_Syntax_SepArray_ofElems(x_190, x_174); +lean_dec_ref(x_174); +lean_inc_ref(x_172); +x_192 = l_Array_append___redArg(x_172, x_191); lean_dec_ref(x_191); -lean_inc(x_175); lean_inc(x_180); +lean_inc(x_179); x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_180); -lean_ctor_set(x_193, 1, x_175); +lean_ctor_set(x_193, 0, x_179); +lean_ctor_set(x_193, 1, x_180); lean_ctor_set(x_193, 2, x_192); x_194 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; -lean_inc(x_180); +lean_inc(x_179); x_195 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_195, 0, x_180); +lean_ctor_set(x_195, 0, x_179); lean_ctor_set(x_195, 1, x_194); +lean_inc(x_179); +x_196 = l_Lean_Syntax_node3(x_179, x_187, x_189, x_193, x_195); lean_inc(x_180); -x_196 = l_Lean_Syntax_node3(x_180, x_187, x_189, x_193, x_195); -lean_inc(x_175); -lean_inc(x_180); -x_197 = l_Lean_Syntax_node1(x_180, x_175, x_196); -lean_inc(x_180); +lean_inc(x_179); +x_197 = l_Lean_Syntax_node1(x_179, x_180, x_196); +lean_inc(x_179); x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_180); -lean_ctor_set(x_198, 1, x_176); +lean_ctor_set(x_198, 0, x_179); +lean_ctor_set(x_198, 1, x_173); x_199 = l_Lean_Elab_Command_elabElabRulesAux___closed__5; x_200 = l_Lean_Elab_Command_elabElabRulesAux___closed__6; -lean_inc(x_173); lean_inc(x_178); -x_201 = l_Lean_addMacroScope(x_178, x_200, x_173); +lean_inc(x_171); +x_201 = l_Lean_addMacroScope(x_171, x_200, x_178); x_202 = lean_box(0); -lean_inc(x_180); +lean_inc(x_179); x_203 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_203, 0, x_180); +lean_ctor_set(x_203, 0, x_179); lean_ctor_set(x_203, 1, x_199); lean_ctor_set(x_203, 2, x_201); lean_ctor_set(x_203, 3, x_202); x_204 = lean_mk_syntax_ident(x_4); -lean_inc(x_175); lean_inc(x_180); -x_205 = l_Lean_Syntax_node2(x_180, x_175, x_203, x_204); +lean_inc(x_179); +x_205 = l_Lean_Syntax_node2(x_179, x_180, x_203, x_204); x_206 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_180); +lean_inc(x_179); x_207 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_207, 0, x_180); +lean_ctor_set(x_207, 0, x_179); lean_ctor_set(x_207, 1, x_206); x_208 = l_Lean_Elab_Command_elabElabRulesAux___closed__9; x_209 = l_Lean_Elab_Command_elabElabRulesAux___closed__10; -lean_inc_ref(x_170); +lean_inc_ref(x_175); lean_inc_ref(x_177); -x_210 = l_Lean_Name_mkStr4(x_177, x_170, x_185, x_209); -lean_inc(x_173); -lean_inc(x_210); +x_210 = l_Lean_Name_mkStr4(x_177, x_175, x_185, x_209); lean_inc(x_178); -x_211 = l_Lean_addMacroScope(x_178, x_210, x_173); +lean_inc(x_210); +lean_inc(x_171); +x_211 = l_Lean_addMacroScope(x_171, x_210, x_178); x_212 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_212, 0, x_210); lean_ctor_set(x_212, 1, x_202); x_213 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_213, 0, x_212); lean_ctor_set(x_213, 1, x_202); -lean_inc(x_180); +lean_inc(x_179); x_214 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_214, 0, x_180); +lean_ctor_set(x_214, 0, x_179); lean_ctor_set(x_214, 1, x_208); lean_ctor_set(x_214, 2, x_211); lean_ctor_set(x_214, 3, x_213); x_215 = l_Lean_Elab_Command_elabElabRulesAux___closed__11; -lean_inc(x_180); +lean_inc(x_179); x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_180); +lean_ctor_set(x_216, 0, x_179); lean_ctor_set(x_216, 1, x_215); x_217 = l_Lean_Elab_Command_elabElabRulesAux___closed__12; lean_inc_ref(x_177); x_218 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_217); -lean_inc(x_180); +lean_inc(x_179); x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_180); +lean_ctor_set(x_219, 0, x_179); lean_ctor_set(x_219, 1, x_217); x_220 = l_Lean_Elab_Command_elabElabRulesAux___closed__13; lean_inc_ref(x_177); x_221 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_220); x_222 = l_Lean_Elab_Command_elabElabRulesAux___closed__15; x_223 = l_Lean_Elab_Command_elabElabRulesAux___closed__16; -lean_inc(x_173); lean_inc(x_178); -x_224 = l_Lean_addMacroScope(x_178, x_223, x_173); -lean_inc(x_180); +lean_inc(x_171); +x_224 = l_Lean_addMacroScope(x_171, x_223, x_178); +lean_inc(x_179); x_225 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_225, 0, x_180); +lean_ctor_set(x_225, 0, x_179); lean_ctor_set(x_225, 1, x_222); lean_ctor_set(x_225, 2, x_224); lean_ctor_set(x_225, 3, x_202); @@ -3504,49 +3504,49 @@ x_226 = l_Lean_Elab_Command_elabElabRulesAux___closed__28; lean_inc_ref(x_177); x_227 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_226); x_228 = l_Lean_Elab_Command_elabElabRulesAux___closed__29; -lean_inc(x_180); +lean_inc(x_179); x_229 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_229, 0, x_180); +lean_ctor_set(x_229, 0, x_179); lean_ctor_set(x_229, 1, x_228); -lean_inc(x_180); -x_230 = l_Lean_Syntax_node1(x_180, x_227, x_229); +lean_inc(x_179); +x_230 = l_Lean_Syntax_node1(x_179, x_227, x_229); lean_inc(x_230); lean_inc_ref(x_225); -lean_inc(x_175); lean_inc(x_180); -x_231 = l_Lean_Syntax_node2(x_180, x_175, x_225, x_230); -lean_inc_ref(x_179); -lean_inc(x_175); +lean_inc(x_179); +x_231 = l_Lean_Syntax_node2(x_179, x_180, x_225, x_230); +lean_inc_ref(x_172); lean_inc(x_180); +lean_inc(x_179); x_232 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_232, 0, x_180); -lean_ctor_set(x_232, 1, x_175); -lean_ctor_set(x_232, 2, x_179); +lean_ctor_set(x_232, 0, x_179); +lean_ctor_set(x_232, 1, x_180); +lean_ctor_set(x_232, 2, x_172); x_233 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; -lean_inc(x_180); +lean_inc(x_179); x_234 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_234, 0, x_180); +lean_ctor_set(x_234, 0, x_179); lean_ctor_set(x_234, 1, x_233); x_235 = l_Lean_Elab_Command_elabElabRulesAux___closed__24; lean_inc_ref(x_177); x_236 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_235); -lean_inc(x_180); +lean_inc(x_179); x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_180); +lean_ctor_set(x_237, 0, x_179); lean_ctor_set(x_237, 1, x_235); x_238 = l_Lean_Elab_Command_elabElabRulesAux___closed__25; lean_inc_ref(x_177); x_239 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_238); lean_inc_ref(x_232); +lean_inc(x_179); +x_240 = l_Lean_Syntax_node2(x_179, x_239, x_232, x_225); lean_inc(x_180); -x_240 = l_Lean_Syntax_node2(x_180, x_239, x_232, x_225); -lean_inc(x_175); -lean_inc(x_180); -x_241 = l_Lean_Syntax_node1(x_180, x_175, x_240); +lean_inc(x_179); +x_241 = l_Lean_Syntax_node1(x_179, x_180, x_240); x_242 = l_Lean_Elab_Command_elabElabRulesAux___closed__26; -lean_inc(x_180); +lean_inc(x_179); x_243 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_243, 0, x_180); +lean_ctor_set(x_243, 0, x_179); lean_ctor_set(x_243, 1, x_242); x_244 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; lean_inc_ref(x_177); @@ -3554,68 +3554,68 @@ x_245 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_244); x_246 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__7; lean_inc_ref(x_177); x_247 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_246); -x_248 = l_Array_append___redArg(x_179, x_14); +x_248 = l_Array_append___redArg(x_172, x_14); lean_dec(x_14); x_249 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__9; -lean_inc(x_180); +lean_inc(x_179); x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_180); +lean_ctor_set(x_250, 0, x_179); lean_ctor_set(x_250, 1, x_249); -lean_inc(x_175); lean_inc(x_180); -x_251 = l_Lean_Syntax_node1(x_180, x_175, x_230); -lean_inc(x_175); +lean_inc(x_179); +x_251 = l_Lean_Syntax_node1(x_179, x_180, x_230); lean_inc(x_180); -x_252 = l_Lean_Syntax_node1(x_180, x_175, x_251); +lean_inc(x_179); +x_252 = l_Lean_Syntax_node1(x_179, x_180, x_251); x_253 = l_Lean_Elab_Command_elabElabRulesAux___closed__30; lean_inc_ref(x_177); x_254 = l_Lean_Name_mkStr4(x_177, x_184, x_185, x_253); x_255 = l_Lean_Elab_Command_elabElabRulesAux___closed__31; -lean_inc(x_180); +lean_inc(x_179); x_256 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_256, 0, x_180); +lean_ctor_set(x_256, 0, x_179); lean_ctor_set(x_256, 1, x_255); x_257 = l_Lean_Elab_Command_elabElabRulesAux___closed__32; x_258 = l_Lean_Elab_Command_elabElabRulesAux___closed__33; x_259 = l_Lean_Elab_Command_elabElabRulesAux___closed__34; -x_260 = l_Lean_addMacroScope(x_178, x_259, x_173); -x_261 = l_Lean_Name_mkStr3(x_177, x_170, x_257); +x_260 = l_Lean_addMacroScope(x_171, x_259, x_178); +x_261 = l_Lean_Name_mkStr3(x_177, x_175, x_257); x_262 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_262, 0, x_261); lean_ctor_set(x_262, 1, x_202); x_263 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_263, 0, x_262); lean_ctor_set(x_263, 1, x_202); -lean_inc(x_180); +lean_inc(x_179); x_264 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_264, 0, x_180); +lean_ctor_set(x_264, 0, x_179); lean_ctor_set(x_264, 1, x_258); lean_ctor_set(x_264, 2, x_260); lean_ctor_set(x_264, 3, x_263); -lean_inc(x_180); -x_265 = l_Lean_Syntax_node2(x_180, x_254, x_256, x_264); +lean_inc(x_179); +x_265 = l_Lean_Syntax_node2(x_179, x_254, x_256, x_264); lean_inc_ref(x_234); -lean_inc(x_180); -x_266 = l_Lean_Syntax_node4(x_180, x_247, x_250, x_252, x_234, x_265); +lean_inc(x_179); +x_266 = l_Lean_Syntax_node4(x_179, x_247, x_250, x_252, x_234, x_265); x_267 = lean_array_push(x_248, x_266); -lean_inc(x_180); +lean_inc(x_179); x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_180); -lean_ctor_set(x_268, 1, x_175); +lean_ctor_set(x_268, 0, x_179); +lean_ctor_set(x_268, 1, x_180); lean_ctor_set(x_268, 2, x_267); -lean_inc(x_180); -x_269 = l_Lean_Syntax_node1(x_180, x_245, x_268); +lean_inc(x_179); +x_269 = l_Lean_Syntax_node1(x_179, x_245, x_268); lean_inc_ref_n(x_232, 2); -lean_inc(x_180); -x_270 = l_Lean_Syntax_node6(x_180, x_236, x_237, x_232, x_232, x_241, x_243, x_269); -lean_inc(x_180); -x_271 = l_Lean_Syntax_node4(x_180, x_221, x_231, x_232, x_234, x_270); -lean_inc(x_180); -x_272 = l_Lean_Syntax_node2(x_180, x_218, x_219, x_271); +lean_inc(x_179); +x_270 = l_Lean_Syntax_node6(x_179, x_236, x_237, x_232, x_232, x_241, x_243, x_269); +lean_inc(x_179); +x_271 = l_Lean_Syntax_node4(x_179, x_221, x_231, x_232, x_234, x_270); +lean_inc(x_179); +x_272 = l_Lean_Syntax_node2(x_179, x_218, x_219, x_271); x_273 = l_Lean_Elab_Command_elabElabRulesAux___closed__35; x_274 = lean_array_push(x_273, x_183); x_275 = lean_array_push(x_274, x_197); -x_276 = lean_array_push(x_275, x_174); +x_276 = lean_array_push(x_275, x_170); x_277 = lean_array_push(x_276, x_198); x_278 = lean_array_push(x_277, x_205); x_279 = lean_array_push(x_278, x_207); @@ -3623,8 +3623,8 @@ x_280 = lean_array_push(x_279, x_214); x_281 = lean_array_push(x_280, x_216); x_282 = lean_array_push(x_281, x_272); x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_180); -lean_ctor_set(x_283, 1, x_169); +lean_ctor_set(x_283, 0, x_179); +lean_ctor_set(x_283, 1, x_176); lean_ctor_set(x_283, 2, x_282); x_284 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_284, 0, x_283); @@ -3643,18 +3643,18 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_298; x_298 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_169 = x_295; -x_170 = x_293; -x_171 = x_286; -x_172 = lean_box(0); -x_173 = x_287; -x_174 = x_288; -x_175 = x_296; -x_176 = x_294; +x_169 = lean_box(0); +x_170 = x_286; +x_171 = x_290; +x_172 = x_297; +x_173 = x_294; +x_174 = x_287; +x_175 = x_293; +x_176 = x_295; x_177 = x_292; -x_178 = x_290; -x_179 = x_297; -x_180 = x_289; +x_178 = x_289; +x_179 = x_288; +x_180 = x_296; x_181 = x_298; goto block_285; } @@ -3665,18 +3665,18 @@ x_299 = lean_ctor_get(x_1, 0); lean_inc(x_299); lean_dec_ref(x_1); x_300 = l_Array_mkArray1___redArg(x_299); -x_169 = x_295; -x_170 = x_293; -x_171 = x_286; -x_172 = lean_box(0); -x_173 = x_287; -x_174 = x_288; -x_175 = x_296; -x_176 = x_294; +x_169 = lean_box(0); +x_170 = x_286; +x_171 = x_290; +x_172 = x_297; +x_173 = x_294; +x_174 = x_287; +x_175 = x_293; +x_176 = x_295; x_177 = x_292; -x_178 = x_290; -x_179 = x_297; -x_180 = x_289; +x_178 = x_289; +x_179 = x_288; +x_180 = x_296; x_181 = x_300; goto block_285; } @@ -3684,182 +3684,182 @@ goto block_285; block_400: { lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; 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_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; -lean_inc_ref(x_311); -x_316 = l_Array_append___redArg(x_311, x_315); +lean_inc_ref(x_314); +x_316 = l_Array_append___redArg(x_314, x_315); lean_dec_ref(x_315); -lean_inc(x_313); -lean_inc(x_309); +lean_inc(x_308); +lean_inc(x_305); x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_309); -lean_ctor_set(x_317, 1, x_313); +lean_ctor_set(x_317, 0, x_305); +lean_ctor_set(x_317, 1, x_308); lean_ctor_set(x_317, 2, x_316); x_318 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__5; x_319 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__6; x_320 = l_Lean_Elab_Command_elabElabRulesAux___closed__0; -lean_inc_ref(x_307); -x_321 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_320); +lean_inc_ref(x_306); +x_321 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_320); x_322 = l_Lean_Elab_Command_elabElabRulesAux___closed__1; -lean_inc(x_309); +lean_inc(x_305); x_323 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_323, 0, x_309); +lean_ctor_set(x_323, 0, x_305); lean_ctor_set(x_323, 1, x_322); x_324 = l_Lean_Elab_Command_elabElabRulesAux___closed__2; -x_325 = l_Lean_Syntax_SepArray_ofElems(x_324, x_305); -lean_dec_ref(x_305); -lean_inc_ref(x_311); -x_326 = l_Array_append___redArg(x_311, x_325); +x_325 = l_Lean_Syntax_SepArray_ofElems(x_324, x_311); +lean_dec_ref(x_311); +lean_inc_ref(x_314); +x_326 = l_Array_append___redArg(x_314, x_325); lean_dec_ref(x_325); -lean_inc(x_313); -lean_inc(x_309); +lean_inc(x_308); +lean_inc(x_305); x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_309); -lean_ctor_set(x_327, 1, x_313); +lean_ctor_set(x_327, 0, x_305); +lean_ctor_set(x_327, 1, x_308); lean_ctor_set(x_327, 2, x_326); x_328 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; -lean_inc(x_309); +lean_inc(x_305); x_329 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_329, 0, x_309); +lean_ctor_set(x_329, 0, x_305); lean_ctor_set(x_329, 1, x_328); -lean_inc(x_309); -x_330 = l_Lean_Syntax_node3(x_309, x_321, x_323, x_327, x_329); -lean_inc(x_313); -lean_inc(x_309); -x_331 = l_Lean_Syntax_node1(x_309, x_313, x_330); -lean_inc(x_309); +lean_inc(x_305); +x_330 = l_Lean_Syntax_node3(x_305, x_321, x_323, x_327, x_329); +lean_inc(x_308); +lean_inc(x_305); +x_331 = l_Lean_Syntax_node1(x_305, x_308, x_330); +lean_inc(x_305); x_332 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_332, 0, x_309); -lean_ctor_set(x_332, 1, x_306); +lean_ctor_set(x_332, 0, x_305); +lean_ctor_set(x_332, 1, x_307); x_333 = l_Lean_Elab_Command_elabElabRulesAux___closed__5; x_334 = l_Lean_Elab_Command_elabElabRulesAux___closed__6; -lean_inc(x_312); +lean_inc(x_304); lean_inc(x_310); -x_335 = l_Lean_addMacroScope(x_310, x_334, x_312); +x_335 = l_Lean_addMacroScope(x_310, x_334, x_304); x_336 = lean_box(0); -lean_inc(x_309); +lean_inc(x_305); x_337 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_337, 0, x_309); +lean_ctor_set(x_337, 0, x_305); lean_ctor_set(x_337, 1, x_333); lean_ctor_set(x_337, 2, x_335); lean_ctor_set(x_337, 3, x_336); x_338 = lean_mk_syntax_ident(x_4); -lean_inc(x_313); -lean_inc(x_309); -x_339 = l_Lean_Syntax_node2(x_309, x_313, x_337, x_338); +lean_inc(x_308); +lean_inc(x_305); +x_339 = l_Lean_Syntax_node2(x_305, x_308, x_337, x_338); x_340 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_309); +lean_inc(x_305); x_341 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_341, 0, x_309); +lean_ctor_set(x_341, 0, x_305); lean_ctor_set(x_341, 1, x_340); x_342 = l_Lean_Elab_Command_elabElabRulesAux___closed__42; x_343 = l_Lean_Elab_Command_elabElabRulesAux___closed__43; -lean_inc_ref(x_314); -lean_inc_ref(x_307); -x_344 = l_Lean_Name_mkStr4(x_307, x_314, x_308, x_343); -lean_inc(x_312); +lean_inc_ref(x_303); +lean_inc_ref(x_306); +x_344 = l_Lean_Name_mkStr4(x_306, x_303, x_309, x_343); +lean_inc(x_304); lean_inc(x_344); lean_inc(x_310); -x_345 = l_Lean_addMacroScope(x_310, x_344, x_312); +x_345 = l_Lean_addMacroScope(x_310, x_344, x_304); x_346 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_346, 0, x_344); lean_ctor_set(x_346, 1, x_336); x_347 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_347, 0, x_346); lean_ctor_set(x_347, 1, x_336); -lean_inc(x_309); +lean_inc(x_305); x_348 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_348, 0, x_309); +lean_ctor_set(x_348, 0, x_305); lean_ctor_set(x_348, 1, x_342); lean_ctor_set(x_348, 2, x_345); lean_ctor_set(x_348, 3, x_347); x_349 = l_Lean_Elab_Command_elabElabRulesAux___closed__11; -lean_inc(x_309); +lean_inc(x_305); x_350 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_350, 0, x_309); +lean_ctor_set(x_350, 0, x_305); lean_ctor_set(x_350, 1, x_349); x_351 = l_Lean_Elab_Command_elabElabRulesAux___closed__12; -lean_inc_ref(x_307); -x_352 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_351); -lean_inc(x_309); +lean_inc_ref(x_306); +x_352 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_351); +lean_inc(x_305); x_353 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_353, 0, x_309); +lean_ctor_set(x_353, 0, x_305); lean_ctor_set(x_353, 1, x_351); x_354 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; -lean_inc_ref(x_307); -x_355 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_354); +lean_inc_ref(x_306); +x_355 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_354); x_356 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__7; -lean_inc_ref(x_307); -x_357 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_356); -x_358 = l_Array_append___redArg(x_311, x_14); +lean_inc_ref(x_306); +x_357 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_356); +x_358 = l_Array_append___redArg(x_314, x_14); lean_dec(x_14); x_359 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__9; -lean_inc(x_309); +lean_inc(x_305); x_360 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_360, 0, x_309); +lean_ctor_set(x_360, 0, x_305); lean_ctor_set(x_360, 1, x_359); x_361 = l_Lean_Elab_Command_elabElabRulesAux___closed__28; -lean_inc_ref(x_307); -x_362 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_361); +lean_inc_ref(x_306); +x_362 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_361); x_363 = l_Lean_Elab_Command_elabElabRulesAux___closed__29; -lean_inc(x_309); +lean_inc(x_305); x_364 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_364, 0, x_309); +lean_ctor_set(x_364, 0, x_305); lean_ctor_set(x_364, 1, x_363); -lean_inc(x_309); -x_365 = l_Lean_Syntax_node1(x_309, x_362, x_364); -lean_inc(x_313); -lean_inc(x_309); -x_366 = l_Lean_Syntax_node1(x_309, x_313, x_365); -lean_inc(x_313); -lean_inc(x_309); -x_367 = l_Lean_Syntax_node1(x_309, x_313, x_366); +lean_inc(x_305); +x_365 = l_Lean_Syntax_node1(x_305, x_362, x_364); +lean_inc(x_308); +lean_inc(x_305); +x_366 = l_Lean_Syntax_node1(x_305, x_308, x_365); +lean_inc(x_308); +lean_inc(x_305); +x_367 = l_Lean_Syntax_node1(x_305, x_308, x_366); x_368 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; -lean_inc(x_309); +lean_inc(x_305); x_369 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_369, 0, x_309); +lean_ctor_set(x_369, 0, x_305); lean_ctor_set(x_369, 1, x_368); x_370 = l_Lean_Elab_Command_elabElabRulesAux___closed__30; -lean_inc_ref(x_307); -x_371 = l_Lean_Name_mkStr4(x_307, x_318, x_319, x_370); +lean_inc_ref(x_306); +x_371 = l_Lean_Name_mkStr4(x_306, x_318, x_319, x_370); x_372 = l_Lean_Elab_Command_elabElabRulesAux___closed__31; -lean_inc(x_309); +lean_inc(x_305); x_373 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_373, 0, x_309); +lean_ctor_set(x_373, 0, x_305); lean_ctor_set(x_373, 1, x_372); x_374 = l_Lean_Elab_Command_elabElabRulesAux___closed__32; x_375 = l_Lean_Elab_Command_elabElabRulesAux___closed__33; x_376 = l_Lean_Elab_Command_elabElabRulesAux___closed__34; -x_377 = l_Lean_addMacroScope(x_310, x_376, x_312); -x_378 = l_Lean_Name_mkStr3(x_307, x_314, x_374); +x_377 = l_Lean_addMacroScope(x_310, x_376, x_304); +x_378 = l_Lean_Name_mkStr3(x_306, x_303, x_374); x_379 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_379, 0, x_378); lean_ctor_set(x_379, 1, x_336); x_380 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_380, 0, x_379); lean_ctor_set(x_380, 1, x_336); -lean_inc(x_309); +lean_inc(x_305); x_381 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_381, 0, x_309); +lean_ctor_set(x_381, 0, x_305); lean_ctor_set(x_381, 1, x_375); lean_ctor_set(x_381, 2, x_377); lean_ctor_set(x_381, 3, x_380); -lean_inc(x_309); -x_382 = l_Lean_Syntax_node2(x_309, x_371, x_373, x_381); -lean_inc(x_309); -x_383 = l_Lean_Syntax_node4(x_309, x_357, x_360, x_367, x_369, x_382); +lean_inc(x_305); +x_382 = l_Lean_Syntax_node2(x_305, x_371, x_373, x_381); +lean_inc(x_305); +x_383 = l_Lean_Syntax_node4(x_305, x_357, x_360, x_367, x_369, x_382); x_384 = lean_array_push(x_358, x_383); -lean_inc(x_309); +lean_inc(x_305); x_385 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_385, 0, x_309); -lean_ctor_set(x_385, 1, x_313); +lean_ctor_set(x_385, 0, x_305); +lean_ctor_set(x_385, 1, x_308); lean_ctor_set(x_385, 2, x_384); -lean_inc(x_309); -x_386 = l_Lean_Syntax_node1(x_309, x_355, x_385); -lean_inc(x_309); -x_387 = l_Lean_Syntax_node2(x_309, x_352, x_353, x_386); +lean_inc(x_305); +x_386 = l_Lean_Syntax_node1(x_305, x_355, x_385); +lean_inc(x_305); +x_387 = l_Lean_Syntax_node2(x_305, x_352, x_353, x_386); x_388 = l_Lean_Elab_Command_elabElabRulesAux___closed__35; x_389 = lean_array_push(x_388, x_317); x_390 = lean_array_push(x_389, x_331); -x_391 = lean_array_push(x_390, x_304); +x_391 = lean_array_push(x_390, x_302); x_392 = lean_array_push(x_391, x_332); x_393 = lean_array_push(x_392, x_339); x_394 = lean_array_push(x_393, x_341); @@ -3867,8 +3867,8 @@ x_395 = lean_array_push(x_394, x_348); x_396 = lean_array_push(x_395, x_350); x_397 = lean_array_push(x_396, x_387); x_398 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_398, 0, x_309); -lean_ctor_set(x_398, 1, x_303); +lean_ctor_set(x_398, 0, x_305); +lean_ctor_set(x_398, 1, x_312); lean_ctor_set(x_398, 2, x_397); x_399 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_399, 0, x_398); @@ -3888,19 +3888,19 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_414; x_414 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_302 = lean_box(0); -x_303 = x_411; -x_304 = x_402; -x_305 = x_403; -x_306 = x_410; -x_307 = x_407; -x_308 = x_409; -x_309 = x_401; +x_302 = x_401; +x_303 = x_408; +x_304 = x_403; +x_305 = x_404; +x_306 = x_407; +x_307 = x_410; +x_308 = x_412; +x_309 = x_409; x_310 = x_405; -x_311 = x_413; -x_312 = x_404; -x_313 = x_412; -x_314 = x_408; +x_311 = x_402; +x_312 = x_411; +x_313 = lean_box(0); +x_314 = x_413; x_315 = x_414; goto block_400; } @@ -3911,19 +3911,19 @@ x_415 = lean_ctor_get(x_1, 0); lean_inc(x_415); lean_dec_ref(x_1); x_416 = l_Array_mkArray1___redArg(x_415); -x_302 = lean_box(0); -x_303 = x_411; -x_304 = x_402; -x_305 = x_403; -x_306 = x_410; -x_307 = x_407; -x_308 = x_409; -x_309 = x_401; +x_302 = x_401; +x_303 = x_408; +x_304 = x_403; +x_305 = x_404; +x_306 = x_407; +x_307 = x_410; +x_308 = x_412; +x_309 = x_409; x_310 = x_405; -x_311 = x_413; -x_312 = x_404; -x_313 = x_412; -x_314 = x_408; +x_311 = x_402; +x_312 = x_411; +x_313 = lean_box(0); +x_314 = x_413; x_315 = x_416; goto block_400; } @@ -3931,182 +3931,182 @@ goto block_400; block_515: { lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; -lean_inc_ref(x_425); -x_431 = l_Array_append___redArg(x_425, x_430); +lean_inc_ref(x_421); +x_431 = l_Array_append___redArg(x_421, x_430); lean_dec_ref(x_430); -lean_inc(x_429); -lean_inc(x_426); +lean_inc(x_425); +lean_inc(x_422); x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_426); -lean_ctor_set(x_432, 1, x_429); +lean_ctor_set(x_432, 0, x_422); +lean_ctor_set(x_432, 1, x_425); lean_ctor_set(x_432, 2, x_431); x_433 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__5; x_434 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__6; x_435 = l_Lean_Elab_Command_elabElabRulesAux___closed__0; -lean_inc_ref(x_423); -x_436 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_435); +lean_inc_ref(x_426); +x_436 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_435); x_437 = l_Lean_Elab_Command_elabElabRulesAux___closed__1; -lean_inc(x_426); +lean_inc(x_422); x_438 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_438, 0, x_426); +lean_ctor_set(x_438, 0, x_422); lean_ctor_set(x_438, 1, x_437); x_439 = l_Lean_Elab_Command_elabElabRulesAux___closed__2; -x_440 = l_Lean_Syntax_SepArray_ofElems(x_439, x_428); -lean_dec_ref(x_428); -lean_inc_ref(x_425); -x_441 = l_Array_append___redArg(x_425, x_440); +x_440 = l_Lean_Syntax_SepArray_ofElems(x_439, x_420); +lean_dec_ref(x_420); +lean_inc_ref(x_421); +x_441 = l_Array_append___redArg(x_421, x_440); lean_dec_ref(x_440); -lean_inc(x_429); -lean_inc(x_426); +lean_inc(x_425); +lean_inc(x_422); x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_426); -lean_ctor_set(x_442, 1, x_429); +lean_ctor_set(x_442, 0, x_422); +lean_ctor_set(x_442, 1, x_425); lean_ctor_set(x_442, 2, x_441); x_443 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; -lean_inc(x_426); +lean_inc(x_422); x_444 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_444, 0, x_426); +lean_ctor_set(x_444, 0, x_422); lean_ctor_set(x_444, 1, x_443); -lean_inc(x_426); -x_445 = l_Lean_Syntax_node3(x_426, x_436, x_438, x_442, x_444); -lean_inc(x_429); -lean_inc(x_426); -x_446 = l_Lean_Syntax_node1(x_426, x_429, x_445); -lean_inc(x_426); +lean_inc(x_422); +x_445 = l_Lean_Syntax_node3(x_422, x_436, x_438, x_442, x_444); +lean_inc(x_425); +lean_inc(x_422); +x_446 = l_Lean_Syntax_node1(x_422, x_425, x_445); +lean_inc(x_422); x_447 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_447, 0, x_426); +lean_ctor_set(x_447, 0, x_422); lean_ctor_set(x_447, 1, x_419); x_448 = l_Lean_Elab_Command_elabElabRulesAux___closed__5; x_449 = l_Lean_Elab_Command_elabElabRulesAux___closed__6; -lean_inc(x_427); -lean_inc(x_424); -x_450 = l_Lean_addMacroScope(x_424, x_449, x_427); +lean_inc(x_428); +lean_inc(x_429); +x_450 = l_Lean_addMacroScope(x_429, x_449, x_428); x_451 = lean_box(0); -lean_inc(x_426); +lean_inc(x_422); x_452 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_452, 0, x_426); +lean_ctor_set(x_452, 0, x_422); lean_ctor_set(x_452, 1, x_448); lean_ctor_set(x_452, 2, x_450); lean_ctor_set(x_452, 3, x_451); x_453 = lean_mk_syntax_ident(x_4); -lean_inc(x_429); -lean_inc(x_426); -x_454 = l_Lean_Syntax_node2(x_426, x_429, x_452, x_453); +lean_inc(x_425); +lean_inc(x_422); +x_454 = l_Lean_Syntax_node2(x_422, x_425, x_452, x_453); x_455 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_426); +lean_inc(x_422); x_456 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_456, 0, x_426); +lean_ctor_set(x_456, 0, x_422); lean_ctor_set(x_456, 1, x_455); x_457 = l_Lean_Elab_Command_elabElabRulesAux___closed__45; x_458 = l_Lean_Elab_Command_elabElabRulesAux___closed__46; -lean_inc_ref(x_421); -lean_inc_ref(x_423); -x_459 = l_Lean_Name_mkStr4(x_423, x_421, x_458, x_458); -lean_inc(x_427); +lean_inc_ref(x_424); +lean_inc_ref(x_426); +x_459 = l_Lean_Name_mkStr4(x_426, x_424, x_458, x_458); +lean_inc(x_428); lean_inc(x_459); -lean_inc(x_424); -x_460 = l_Lean_addMacroScope(x_424, x_459, x_427); +lean_inc(x_429); +x_460 = l_Lean_addMacroScope(x_429, x_459, x_428); x_461 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_461, 0, x_459); lean_ctor_set(x_461, 1, x_451); x_462 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_462, 0, x_461); lean_ctor_set(x_462, 1, x_451); -lean_inc(x_426); +lean_inc(x_422); x_463 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_463, 0, x_426); +lean_ctor_set(x_463, 0, x_422); lean_ctor_set(x_463, 1, x_457); lean_ctor_set(x_463, 2, x_460); lean_ctor_set(x_463, 3, x_462); x_464 = l_Lean_Elab_Command_elabElabRulesAux___closed__11; -lean_inc(x_426); +lean_inc(x_422); x_465 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_465, 0, x_426); +lean_ctor_set(x_465, 0, x_422); lean_ctor_set(x_465, 1, x_464); x_466 = l_Lean_Elab_Command_elabElabRulesAux___closed__12; -lean_inc_ref(x_423); -x_467 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_466); -lean_inc(x_426); +lean_inc_ref(x_426); +x_467 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_466); +lean_inc(x_422); x_468 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_468, 0, x_426); +lean_ctor_set(x_468, 0, x_422); lean_ctor_set(x_468, 1, x_466); x_469 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; -lean_inc_ref(x_423); -x_470 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_469); +lean_inc_ref(x_426); +x_470 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_469); x_471 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__7; -lean_inc_ref(x_423); -x_472 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_471); -x_473 = l_Array_append___redArg(x_425, x_14); +lean_inc_ref(x_426); +x_472 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_471); +x_473 = l_Array_append___redArg(x_421, x_14); lean_dec(x_14); x_474 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__9; -lean_inc(x_426); +lean_inc(x_422); x_475 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_475, 0, x_426); +lean_ctor_set(x_475, 0, x_422); lean_ctor_set(x_475, 1, x_474); x_476 = l_Lean_Elab_Command_elabElabRulesAux___closed__28; -lean_inc_ref(x_423); -x_477 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_476); +lean_inc_ref(x_426); +x_477 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_476); x_478 = l_Lean_Elab_Command_elabElabRulesAux___closed__29; -lean_inc(x_426); +lean_inc(x_422); x_479 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_479, 0, x_426); +lean_ctor_set(x_479, 0, x_422); lean_ctor_set(x_479, 1, x_478); -lean_inc(x_426); -x_480 = l_Lean_Syntax_node1(x_426, x_477, x_479); -lean_inc(x_429); -lean_inc(x_426); -x_481 = l_Lean_Syntax_node1(x_426, x_429, x_480); -lean_inc(x_429); -lean_inc(x_426); -x_482 = l_Lean_Syntax_node1(x_426, x_429, x_481); +lean_inc(x_422); +x_480 = l_Lean_Syntax_node1(x_422, x_477, x_479); +lean_inc(x_425); +lean_inc(x_422); +x_481 = l_Lean_Syntax_node1(x_422, x_425, x_480); +lean_inc(x_425); +lean_inc(x_422); +x_482 = l_Lean_Syntax_node1(x_422, x_425, x_481); x_483 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; -lean_inc(x_426); +lean_inc(x_422); x_484 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_484, 0, x_426); +lean_ctor_set(x_484, 0, x_422); lean_ctor_set(x_484, 1, x_483); x_485 = l_Lean_Elab_Command_elabElabRulesAux___closed__30; -lean_inc_ref(x_423); -x_486 = l_Lean_Name_mkStr4(x_423, x_433, x_434, x_485); +lean_inc_ref(x_426); +x_486 = l_Lean_Name_mkStr4(x_426, x_433, x_434, x_485); x_487 = l_Lean_Elab_Command_elabElabRulesAux___closed__31; -lean_inc(x_426); +lean_inc(x_422); x_488 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_488, 0, x_426); +lean_ctor_set(x_488, 0, x_422); lean_ctor_set(x_488, 1, x_487); x_489 = l_Lean_Elab_Command_elabElabRulesAux___closed__32; x_490 = l_Lean_Elab_Command_elabElabRulesAux___closed__33; x_491 = l_Lean_Elab_Command_elabElabRulesAux___closed__34; -x_492 = l_Lean_addMacroScope(x_424, x_491, x_427); -x_493 = l_Lean_Name_mkStr3(x_423, x_421, x_489); +x_492 = l_Lean_addMacroScope(x_429, x_491, x_428); +x_493 = l_Lean_Name_mkStr3(x_426, x_424, x_489); x_494 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_494, 0, x_493); lean_ctor_set(x_494, 1, x_451); x_495 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_495, 0, x_494); lean_ctor_set(x_495, 1, x_451); -lean_inc(x_426); +lean_inc(x_422); x_496 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_496, 0, x_426); +lean_ctor_set(x_496, 0, x_422); lean_ctor_set(x_496, 1, x_490); lean_ctor_set(x_496, 2, x_492); lean_ctor_set(x_496, 3, x_495); -lean_inc(x_426); -x_497 = l_Lean_Syntax_node2(x_426, x_486, x_488, x_496); -lean_inc(x_426); -x_498 = l_Lean_Syntax_node4(x_426, x_472, x_475, x_482, x_484, x_497); +lean_inc(x_422); +x_497 = l_Lean_Syntax_node2(x_422, x_486, x_488, x_496); +lean_inc(x_422); +x_498 = l_Lean_Syntax_node4(x_422, x_472, x_475, x_482, x_484, x_497); x_499 = lean_array_push(x_473, x_498); -lean_inc(x_426); +lean_inc(x_422); x_500 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_500, 0, x_426); -lean_ctor_set(x_500, 1, x_429); +lean_ctor_set(x_500, 0, x_422); +lean_ctor_set(x_500, 1, x_425); lean_ctor_set(x_500, 2, x_499); -lean_inc(x_426); -x_501 = l_Lean_Syntax_node1(x_426, x_470, x_500); -lean_inc(x_426); -x_502 = l_Lean_Syntax_node2(x_426, x_467, x_468, x_501); +lean_inc(x_422); +x_501 = l_Lean_Syntax_node1(x_422, x_470, x_500); +lean_inc(x_422); +x_502 = l_Lean_Syntax_node2(x_422, x_467, x_468, x_501); x_503 = l_Lean_Elab_Command_elabElabRulesAux___closed__35; x_504 = lean_array_push(x_503, x_432); x_505 = lean_array_push(x_504, x_446); -x_506 = lean_array_push(x_505, x_422); +x_506 = lean_array_push(x_505, x_418); x_507 = lean_array_push(x_506, x_447); x_508 = lean_array_push(x_507, x_454); x_509 = lean_array_push(x_508, x_456); @@ -4114,8 +4114,8 @@ x_510 = lean_array_push(x_509, x_463); x_511 = lean_array_push(x_510, x_465); x_512 = lean_array_push(x_511, x_502); x_513 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_513, 0, x_426); -lean_ctor_set(x_513, 1, x_418); +lean_ctor_set(x_513, 0, x_422); +lean_ctor_set(x_513, 1, x_423); lean_ctor_set(x_513, 2, x_512); x_514 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_514, 0, x_513); @@ -4134,18 +4134,18 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_528; x_528 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_418 = x_525; +x_418 = x_516; x_419 = x_524; -x_420 = lean_box(0); -x_421 = x_523; -x_422 = x_516; -x_423 = x_522; -x_424 = x_520; -x_425 = x_527; -x_426 = x_517; -x_427 = x_519; -x_428 = x_518; -x_429 = x_526; +x_420 = x_517; +x_421 = x_527; +x_422 = x_518; +x_423 = x_525; +x_424 = x_523; +x_425 = x_526; +x_426 = x_522; +x_427 = lean_box(0); +x_428 = x_519; +x_429 = x_520; x_430 = x_528; goto block_515; } @@ -4156,18 +4156,18 @@ x_529 = lean_ctor_get(x_1, 0); lean_inc(x_529); lean_dec_ref(x_1); x_530 = l_Array_mkArray1___redArg(x_529); -x_418 = x_525; +x_418 = x_516; x_419 = x_524; -x_420 = lean_box(0); -x_421 = x_523; -x_422 = x_516; -x_423 = x_522; -x_424 = x_520; -x_425 = x_527; -x_426 = x_517; -x_427 = x_519; -x_428 = x_518; -x_429 = x_526; +x_420 = x_517; +x_421 = x_527; +x_422 = x_518; +x_423 = x_525; +x_424 = x_523; +x_425 = x_526; +x_426 = x_522; +x_427 = lean_box(0); +x_428 = x_519; +x_429 = x_520; x_430 = x_530; goto block_515; } @@ -4176,7 +4176,7 @@ goto block_515; { lean_object* x_538; lean_inc(x_4); -x_538 = l_Lean_Elab_Command_elabElabRulesAux___lam__0(x_4, x_3, x_2, x_537, x_535, x_533); +x_538 = l_Lean_Elab_Command_elabElabRulesAux___lam__0(x_4, x_3, x_2, x_533, x_535, x_534); if (lean_obj_tag(x_538) == 0) { lean_object* x_539; lean_object* x_540; @@ -4200,18 +4200,18 @@ lean_dec_ref(x_542); x_544 = lean_ctor_get(x_535, 5); lean_inc(x_544); lean_dec_ref(x_535); -x_545 = l_Lean_SourceInfo_fromRef(x_541, x_532); +x_545 = l_Lean_SourceInfo_fromRef(x_541, x_536); lean_dec(x_541); if (lean_obj_tag(x_544) == 0) { lean_object* x_546; lean_object* x_547; -x_546 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_533); +x_546 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_534); x_547 = lean_ctor_get(x_546, 0); lean_inc(x_547); lean_dec_ref(x_546); -x_516 = x_534; -x_517 = x_545; -x_518 = x_539; +x_516 = x_532; +x_517 = x_539; +x_518 = x_545; x_519 = x_543; x_520 = x_547; x_521 = lean_box(0); @@ -4223,9 +4223,9 @@ lean_object* x_548; x_548 = lean_ctor_get(x_544, 0); lean_inc(x_548); lean_dec_ref(x_544); -x_516 = x_534; -x_517 = x_545; -x_518 = x_539; +x_516 = x_532; +x_517 = x_539; +x_518 = x_545; x_519 = x_543; x_520 = x_548; x_521 = lean_box(0); @@ -4238,7 +4238,7 @@ uint8_t x_549; lean_dec(x_541); lean_dec(x_539); lean_dec_ref(x_535); -lean_dec(x_534); +lean_dec(x_532); lean_dec(x_14); lean_dec(x_4); lean_dec(x_1); @@ -4263,7 +4263,7 @@ else { lean_dec(x_539); lean_dec_ref(x_535); -lean_dec(x_534); +lean_dec(x_532); lean_dec(x_14); lean_dec(x_4); lean_dec(x_1); @@ -4274,7 +4274,7 @@ else { uint8_t x_552; lean_dec_ref(x_535); -lean_dec(x_534); +lean_dec(x_532); lean_dec(x_14); lean_dec(x_4); lean_dec(x_1); @@ -4344,24 +4344,24 @@ return x_574; else { lean_dec(x_556); -x_532 = x_564; -x_533 = x_558; -x_534 = x_560; +x_532 = x_560; +x_533 = x_565; +x_534 = x_558; x_535 = x_557; -x_536 = lean_box(0); -x_537 = x_565; +x_536 = x_564; +x_537 = lean_box(0); goto block_555; } } else { lean_dec(x_556); -x_532 = x_564; -x_533 = x_558; -x_534 = x_560; +x_532 = x_560; +x_533 = x_565; +x_534 = x_558; x_535 = x_557; -x_536 = lean_box(0); -x_537 = x_565; +x_536 = x_564; +x_537 = lean_box(0); goto block_555; } } @@ -4404,10 +4404,10 @@ x_584 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1 x_585 = lean_ctor_get(x_584, 0); lean_inc(x_585); lean_dec_ref(x_584); -x_401 = x_583; -x_402 = x_560; -x_403 = x_577; -x_404 = x_581; +x_401 = x_560; +x_402 = x_577; +x_403 = x_581; +x_404 = x_583; x_405 = x_585; x_406 = lean_box(0); goto block_417; @@ -4418,10 +4418,10 @@ lean_object* x_586; x_586 = lean_ctor_get(x_582, 0); lean_inc(x_586); lean_dec_ref(x_582); -x_401 = x_583; -x_402 = x_560; -x_403 = x_577; -x_404 = x_581; +x_401 = x_560; +x_402 = x_577; +x_403 = x_581; +x_404 = x_583; x_405 = x_586; x_406 = lean_box(0); goto block_417; @@ -4531,10 +4531,10 @@ x_603 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1 x_604 = lean_ctor_get(x_603, 0); lean_inc(x_604); lean_dec_ref(x_603); -x_286 = x_595; -x_287 = x_599; -x_288 = x_560; -x_289 = x_602; +x_286 = x_560; +x_287 = x_595; +x_288 = x_602; +x_289 = x_599; x_290 = x_604; x_291 = lean_box(0); goto block_301; @@ -4545,10 +4545,10 @@ lean_object* x_605; x_605 = lean_ctor_get(x_600, 0); lean_inc(x_605); lean_dec_ref(x_600); -x_286 = x_595; -x_287 = x_599; -x_288 = x_560; -x_289 = x_602; +x_286 = x_560; +x_287 = x_595; +x_288 = x_602; +x_289 = x_599; x_290 = x_605; x_291 = lean_box(0); goto block_301; @@ -4688,11 +4688,11 @@ x_631 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1 x_632 = lean_ctor_get(x_631, 0); lean_inc(x_632); lean_dec_ref(x_631); -x_152 = x_612; +x_152 = x_560; x_153 = x_627; -x_154 = x_560; +x_154 = x_630; x_155 = x_623; -x_156 = x_630; +x_156 = x_612; x_157 = x_632; x_158 = lean_box(0); goto block_168; @@ -4703,11 +4703,11 @@ lean_object* x_633; x_633 = lean_ctor_get(x_628, 0); lean_inc(x_633); lean_dec_ref(x_628); -x_152 = x_612; +x_152 = x_560; x_153 = x_627; -x_154 = x_560; +x_154 = x_630; x_155 = x_623; -x_156 = x_630; +x_156 = x_612; x_157 = x_633; x_158 = lean_box(0); goto block_168; @@ -5093,24 +5093,24 @@ goto block_118; block_40: { 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_inc_ref(x_26); -x_33 = l_Array_append___redArg(x_26, x_32); +lean_inc_ref(x_27); +x_33 = l_Array_append___redArg(x_27, x_32); lean_dec_ref(x_32); -lean_inc(x_31); +lean_inc(x_28); lean_inc(x_23); x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_23); -lean_ctor_set(x_34, 1, x_31); +lean_ctor_set(x_34, 1, x_28); lean_ctor_set(x_34, 2, x_33); -x_35 = l_Array_append___redArg(x_26, x_13); +x_35 = l_Array_append___redArg(x_27, x_13); lean_inc(x_23); x_36 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_36, 0, x_23); -lean_ctor_set(x_36, 1, x_31); +lean_ctor_set(x_36, 1, x_28); lean_ctor_set(x_36, 2, x_35); lean_inc(x_23); x_37 = l_Lean_Syntax_node1(x_23, x_1, x_36); -x_38 = l_Lean_Syntax_node8(x_23, x_2, x_29, x_28, x_3, x_27, x_30, x_24, x_34, x_37); +x_38 = l_Lean_Syntax_node8(x_23, x_2, x_31, x_30, x_3, x_24, x_25, x_26, x_34, x_37); if (lean_is_scalar(x_20)) { x_39 = lean_alloc_ctor(0, 1, 0); } else { @@ -5122,25 +5122,25 @@ return x_39; block_56: { lean_object* x_49; lean_object* x_50; -lean_inc_ref(x_42); -x_49 = l_Array_append___redArg(x_42, x_48); +lean_inc_ref(x_43); +x_49 = l_Array_append___redArg(x_43, x_48); lean_dec_ref(x_48); -lean_inc(x_47); +lean_inc(x_44); lean_inc(x_23); x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_23); -lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 1, x_44); lean_ctor_set(x_50, 2, x_49); if (lean_obj_tag(x_4) == 0) { lean_object* x_51; x_51 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_24 = x_50; -x_25 = lean_box(0); -x_26 = x_42; +x_24 = x_41; +x_25 = x_42; +x_26 = x_50; x_27 = x_43; -x_28 = x_45; -x_29 = x_44; +x_28 = x_44; +x_29 = lean_box(0); x_30 = x_46; x_31 = x_47; x_32 = x_51; @@ -5158,12 +5158,12 @@ x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_23); lean_ctor_set(x_54, 1, x_53); x_55 = l_Array_mkArray2___redArg(x_54, x_52); -x_24 = x_50; -x_25 = lean_box(0); -x_26 = x_42; +x_24 = x_41; +x_25 = x_42; +x_26 = x_50; x_27 = x_43; -x_28 = x_45; -x_29 = x_44; +x_28 = x_44; +x_29 = lean_box(0); x_30 = x_46; x_31 = x_47; x_32 = x_55; @@ -5176,22 +5176,22 @@ lean_object* x_64; lean_object* x_65; lean_inc_ref(x_58); x_64 = l_Array_append___redArg(x_58, x_63); lean_dec_ref(x_63); -lean_inc(x_62); +lean_inc(x_59); lean_inc(x_23); x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_23); -lean_ctor_set(x_65, 1, x_62); +lean_ctor_set(x_65, 1, x_59); lean_ctor_set(x_65, 2, x_64); if (lean_obj_tag(x_5) == 0) { lean_object* x_66; x_66 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_41 = lean_box(0); -x_42 = x_58; -x_43 = x_59; -x_44 = x_61; -x_45 = x_60; -x_46 = x_65; +x_41 = x_57; +x_42 = x_65; +x_43 = x_58; +x_44 = x_59; +x_45 = lean_box(0); +x_46 = x_61; x_47 = x_62; x_48 = x_66; goto block_56; @@ -5208,12 +5208,12 @@ x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_23); lean_ctor_set(x_69, 1, x_68); x_70 = l_Array_mkArray2___redArg(x_69, x_67); -x_41 = lean_box(0); -x_42 = x_58; -x_43 = x_59; -x_44 = x_61; -x_45 = x_60; -x_46 = x_65; +x_41 = x_57; +x_42 = x_65; +x_43 = x_58; +x_44 = x_59; +x_45 = lean_box(0); +x_46 = x_61; x_47 = x_62; x_48 = x_70; goto block_56; @@ -5222,14 +5222,14 @@ goto block_56; block_92: { lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_inc_ref(x_73); -x_77 = l_Array_append___redArg(x_73, x_76); +lean_inc_ref(x_72); +x_77 = l_Array_append___redArg(x_72, x_76); lean_dec_ref(x_76); -lean_inc(x_75); +lean_inc(x_73); lean_inc(x_23); x_78 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_78, 0, x_23); -lean_ctor_set(x_78, 1, x_75); +lean_ctor_set(x_78, 1, x_73); lean_ctor_set(x_78, 2, x_77); lean_inc(x_23); x_79 = lean_alloc_ctor(2, 2, 0); @@ -5239,11 +5239,11 @@ if (lean_obj_tag(x_12) == 0) { lean_object* x_80; x_80 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_57 = lean_box(0); -x_58 = x_73; -x_59 = x_79; -x_60 = x_78; -x_61 = x_74; +x_57 = x_79; +x_58 = x_72; +x_59 = x_73; +x_60 = lean_box(0); +x_61 = x_78; x_62 = x_75; x_63 = x_80; goto block_71; @@ -5276,11 +5276,11 @@ x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_23); lean_ctor_set(x_90, 1, x_89); x_91 = l_Array_mkArray5___redArg(x_84, x_86, x_88, x_82, x_90); -x_57 = lean_box(0); -x_58 = x_73; -x_59 = x_79; -x_60 = x_78; -x_61 = x_74; +x_57 = x_79; +x_58 = x_72; +x_59 = x_73; +x_60 = lean_box(0); +x_61 = x_78; x_62 = x_75; x_63 = x_91; goto block_71; @@ -5289,14 +5289,14 @@ goto block_71; block_111: { lean_object* x_97; lean_object* x_98; -lean_inc_ref(x_94); -x_97 = l_Array_append___redArg(x_94, x_96); +lean_inc_ref(x_93); +x_97 = l_Array_append___redArg(x_93, x_96); lean_dec_ref(x_96); -lean_inc(x_95); +lean_inc(x_94); lean_inc(x_23); x_98 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_98, 0, x_23); -lean_ctor_set(x_98, 1, x_95); +lean_ctor_set(x_98, 1, x_94); lean_ctor_set(x_98, 2, x_97); if (lean_obj_tag(x_7) == 0) { @@ -5305,10 +5305,10 @@ lean_dec_ref(x_10); lean_dec_ref(x_9); lean_dec_ref(x_8); x_99 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_72 = lean_box(0); +x_72 = x_93; x_73 = x_94; -x_74 = x_98; -x_75 = x_95; +x_74 = lean_box(0); +x_75 = x_98; x_76 = x_99; goto block_92; } @@ -5323,13 +5323,13 @@ lean_inc(x_23); x_104 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_104, 0, x_23); lean_ctor_set(x_104, 1, x_103); -lean_inc_ref(x_94); -x_105 = l_Array_append___redArg(x_94, x_100); -lean_inc(x_95); +lean_inc_ref(x_93); +x_105 = l_Array_append___redArg(x_93, x_100); +lean_inc(x_94); lean_inc(x_23); x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_23); -lean_ctor_set(x_106, 1, x_95); +lean_ctor_set(x_106, 1, x_94); lean_ctor_set(x_106, 2, x_105); x_107 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; lean_inc(x_23); @@ -5339,10 +5339,10 @@ lean_ctor_set(x_108, 1, x_107); lean_inc(x_23); x_109 = l_Lean_Syntax_node3(x_23, x_102, x_104, x_106, x_108); x_110 = l_Array_mkArray1___redArg(x_109); -x_72 = lean_box(0); +x_72 = x_93; x_73 = x_94; -x_74 = x_98; -x_75 = x_95; +x_74 = lean_box(0); +x_75 = x_98; x_76 = x_110; goto block_92; } @@ -5356,9 +5356,9 @@ if (lean_obj_tag(x_11) == 0) { lean_object* x_115; x_115 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_93 = lean_box(0); -x_94 = x_114; -x_95 = x_113; +x_93 = x_114; +x_94 = x_113; +x_95 = lean_box(0); x_96 = x_115; goto block_111; } @@ -5369,9 +5369,9 @@ x_116 = lean_ctor_get(x_11, 0); lean_inc(x_116); lean_dec_ref(x_11); x_117 = l_Array_mkArray1___redArg(x_116); -x_93 = lean_box(0); -x_94 = x_114; -x_95 = x_113; +x_93 = x_114; +x_94 = x_113; +x_95 = lean_box(0); x_96 = x_117; goto block_111; } @@ -5623,11 +5623,11 @@ if (x_24 == 0) lean_object* x_25; lean_dec(x_21); lean_dec(x_19); -lean_dec(x_18); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec_ref(x_13); lean_dec(x_12); x_25 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_25; @@ -5638,20 +5638,20 @@ lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabElabRules___lam__0___boxed), 16, 11); lean_closure_set(x_26, 0, x_23); lean_closure_set(x_26, 1, x_8); -lean_closure_set(x_26, 2, x_15); +lean_closure_set(x_26, 2, x_17); lean_closure_set(x_26, 3, x_19); -lean_closure_set(x_26, 4, x_12); +lean_closure_set(x_26, 4, x_16); lean_closure_set(x_26, 5, x_7); -lean_closure_set(x_26, 6, x_16); +lean_closure_set(x_26, 6, x_15); lean_closure_set(x_26, 7, x_5); lean_closure_set(x_26, 8, x_6); lean_closure_set(x_26, 9, x_22); -lean_closure_set(x_26, 10, x_14); +lean_closure_set(x_26, 10, x_12); x_27 = l_Lean_Syntax_getArg(x_21, x_11); lean_dec(x_21); x_28 = l_Lean_Syntax_getArgs(x_27); lean_dec(x_27); -x_29 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_28, x_7, x_26, x_13, x_18); +x_29 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_28, x_7, x_26, x_18, x_14); lean_dec_ref(x_28); if (lean_obj_tag(x_29) == 0) { @@ -5703,7 +5703,7 @@ if (x_48 == 0) { uint8_t x_49; lean_inc(x_47); -x_49 = l_Lean_Syntax_matchesNull(x_47, x_41); +x_49 = l_Lean_Syntax_matchesNull(x_47, x_40); if (x_49 == 0) { lean_object* x_50; @@ -5721,17 +5721,17 @@ return x_50; else { lean_object* x_51; lean_object* x_52; -x_51 = l_Lean_Syntax_getArg(x_47, x_40); +x_51 = l_Lean_Syntax_getArg(x_47, x_41); lean_dec(x_47); x_52 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_52, 0, x_51); -x_12 = x_42; -x_13 = x_43; -x_14 = x_38; -x_15 = x_37; -x_16 = x_39; -x_17 = lean_box(0); -x_18 = x_44; +x_12 = x_37; +x_13 = lean_box(0); +x_14 = x_44; +x_15 = x_38; +x_16 = x_42; +x_17 = x_39; +x_18 = x_43; x_19 = x_52; goto block_36; } @@ -5741,13 +5741,13 @@ else lean_object* x_53; lean_dec(x_47); x_53 = lean_box(0); -x_12 = x_42; -x_13 = x_43; -x_14 = x_38; -x_15 = x_37; -x_16 = x_39; -x_17 = lean_box(0); -x_18 = x_44; +x_12 = x_37; +x_13 = lean_box(0); +x_14 = x_44; +x_15 = x_38; +x_16 = x_42; +x_17 = x_39; +x_18 = x_43; x_19 = x_53; goto block_36; } @@ -5759,7 +5759,7 @@ x_65 = lean_unsigned_to_nat(7u); x_66 = l_Lean_Syntax_getArg(x_1, x_65); lean_dec(x_1); x_67 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; -x_68 = l_Lean_Name_mkStr4(x_5, x_6, x_58, x_67); +x_68 = l_Lean_Name_mkStr4(x_5, x_6, x_55, x_67); lean_inc(x_66); x_69 = l_Lean_Syntax_isOfKind(x_66, x_68); lean_dec(x_68); @@ -5772,17 +5772,17 @@ lean_dec_ref(x_62); lean_dec(x_61); lean_dec(x_60); lean_dec(x_59); +lean_dec(x_58); lean_dec(x_57); lean_dec(x_56); -lean_dec(x_55); x_70 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_70; } else { lean_object* x_71; lean_object* x_72; -x_71 = l_Lean_TSyntax_getId(x_59); -lean_dec(x_59); +x_71 = l_Lean_TSyntax_getId(x_56); +lean_dec(x_56); lean_inc_ref(x_62); x_72 = l_Lean_Elab_Command_resolveSyntaxKind(x_71, x_62, x_63); if (lean_obj_tag(x_72) == 0) @@ -5795,10 +5795,10 @@ x_74 = l_Lean_Syntax_getArg(x_66, x_11); lean_dec(x_66); x_75 = l_Lean_Syntax_getArgs(x_74); lean_dec(x_74); -x_76 = l_Lean_Elab_Command_elabElabRulesAux(x_56, x_60, x_55, x_73, x_57, x_61, x_75, x_62, x_63); +x_76 = l_Lean_Elab_Command_elabElabRulesAux(x_60, x_59, x_57, x_73, x_58, x_61, x_75, x_62, x_63); lean_dec(x_63); -lean_dec(x_57); -lean_dec(x_60); +lean_dec(x_58); +lean_dec(x_59); return x_76; } else @@ -5809,9 +5809,9 @@ lean_dec(x_63); lean_dec_ref(x_62); lean_dec(x_61); lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_58); lean_dec(x_57); -lean_dec(x_56); -lean_dec(x_55); x_77 = !lean_is_exclusive(x_72); if (x_77 == 0) { @@ -5840,7 +5840,7 @@ if (x_94 == 0) { uint8_t x_95; lean_inc(x_93); -x_95 = l_Lean_Syntax_matchesNull(x_93, x_87); +x_95 = l_Lean_Syntax_matchesNull(x_93, x_81); if (x_95 == 0) { lean_object* x_96; @@ -5848,11 +5848,11 @@ lean_dec(x_93); lean_dec(x_90); lean_dec_ref(x_89); lean_dec(x_88); +lean_dec(x_87); lean_dec(x_86); lean_dec(x_85); -lean_dec_ref(x_83); -lean_dec(x_82); -lean_dec(x_81); +lean_dec(x_84); +lean_dec_ref(x_82); lean_dec(x_1); x_96 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_96; @@ -5860,15 +5860,15 @@ return x_96; else { lean_object* x_97; lean_object* x_98; -x_97 = l_Lean_Syntax_getArg(x_93, x_84); +x_97 = l_Lean_Syntax_getArg(x_93, x_83); lean_dec(x_93); x_98 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_98, 0, x_97); -x_55 = x_81; -x_56 = x_82; -x_57 = x_88; -x_58 = x_83; -x_59 = x_85; +x_55 = x_82; +x_56 = x_84; +x_57 = x_85; +x_58 = x_88; +x_59 = x_87; x_60 = x_86; x_61 = x_98; x_62 = x_89; @@ -5882,11 +5882,11 @@ else lean_object* x_99; lean_dec(x_93); x_99 = lean_box(0); -x_55 = x_81; -x_56 = x_82; -x_57 = x_88; -x_58 = x_83; -x_59 = x_85; +x_55 = x_82; +x_56 = x_84; +x_57 = x_85; +x_58 = x_88; +x_59 = x_87; x_60 = x_86; x_61 = x_99; x_62 = x_89; @@ -5910,8 +5910,8 @@ lean_object* x_112; lean_dec(x_108); lean_dec(x_106); lean_dec(x_104); -lean_dec(x_103); lean_dec_ref(x_102); +lean_dec(x_101); lean_dec(x_1); x_112 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_112; @@ -5936,8 +5936,8 @@ lean_dec(x_114); lean_dec(x_108); lean_dec(x_106); lean_dec(x_104); -lean_dec(x_103); lean_dec_ref(x_102); +lean_dec(x_101); lean_dec(x_1); x_118 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_118; @@ -5963,8 +5963,8 @@ lean_dec(x_120); lean_dec(x_108); lean_dec(x_106); lean_dec(x_104); -lean_dec(x_103); lean_dec_ref(x_102); +lean_dec(x_101); lean_dec(x_1); x_124 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_124; @@ -5976,16 +5976,16 @@ x_125 = l_Lean_Syntax_getArg(x_121, x_105); lean_dec(x_121); x_126 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_126, 0, x_125); -x_81 = x_108; -x_82 = x_104; -x_83 = x_109; -x_84 = x_105; -x_85 = x_120; -x_86 = x_106; -x_87 = x_107; +x_81 = x_107; +x_82 = x_109; +x_83 = x_105; +x_84 = x_120; +x_85 = x_108; +x_86 = x_101; +x_87 = x_106; x_88 = x_126; x_89 = x_102; -x_90 = x_103; +x_90 = x_104; x_91 = lean_box(0); goto block_100; } @@ -5995,16 +5995,16 @@ else lean_object* x_127; lean_dec(x_121); x_127 = lean_box(0); -x_81 = x_108; -x_82 = x_104; -x_83 = x_109; -x_84 = x_105; -x_85 = x_120; -x_86 = x_106; -x_87 = x_107; +x_81 = x_107; +x_82 = x_109; +x_83 = x_105; +x_84 = x_120; +x_85 = x_108; +x_86 = x_101; +x_87 = x_106; x_88 = x_127; x_89 = x_102; -x_90 = x_103; +x_90 = x_104; x_91 = lean_box(0); goto block_100; } @@ -6029,8 +6029,8 @@ lean_dec(x_129); lean_dec(x_108); lean_dec(x_106); lean_dec(x_104); -lean_dec(x_103); lean_dec_ref(x_102); +lean_dec(x_101); lean_dec(x_1); x_132 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_132; @@ -6042,14 +6042,14 @@ x_133 = l_Lean_Syntax_getArg(x_129, x_105); lean_dec(x_129); x_134 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_134, 0, x_133); -x_37 = x_108; -x_38 = x_104; -x_39 = x_106; -x_40 = x_105; -x_41 = x_107; +x_37 = x_101; +x_38 = x_106; +x_39 = x_108; +x_40 = x_107; +x_41 = x_105; x_42 = x_134; x_43 = x_102; -x_44 = x_103; +x_44 = x_104; x_45 = lean_box(0); goto block_54; } @@ -6059,14 +6059,14 @@ else lean_object* x_135; lean_dec(x_129); x_135 = lean_box(0); -x_37 = x_108; -x_38 = x_104; -x_39 = x_106; -x_40 = x_105; -x_41 = x_107; +x_37 = x_101; +x_38 = x_106; +x_39 = x_108; +x_40 = x_107; +x_41 = x_105; x_42 = x_135; x_43 = x_102; -x_44 = x_103; +x_44 = x_104; x_45 = lean_box(0); goto block_54; } @@ -6123,10 +6123,10 @@ x_151 = l_Lean_Syntax_getArgs(x_150); lean_dec(x_150); x_152 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_152, 0, x_151); -x_101 = lean_box(0); +x_101 = x_137; x_102 = x_138; -x_103 = x_139; -x_104 = x_137; +x_103 = lean_box(0); +x_104 = x_139; x_105 = x_141; x_106 = x_152; goto block_136; @@ -6138,10 +6138,10 @@ else lean_object* x_153; lean_dec(x_142); x_153 = lean_box(0); -x_101 = lean_box(0); +x_101 = x_137; x_102 = x_138; -x_103 = x_139; -x_104 = x_137; +x_103 = lean_box(0); +x_104 = x_139; x_105 = x_141; x_106 = x_153; goto block_136; @@ -8566,7 +8566,7 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab(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; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; 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_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_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; size_t x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; size_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_199; lean_object* x_216; uint8_t 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; uint8_t x_226; lean_object* x_227; lean_object* x_228; size_t x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_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; size_t x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; 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; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; size_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; 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_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_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; size_t x_117; uint8_t x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; 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; size_t x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_216; uint8_t x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; size_t x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; size_t x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; size_t x_309; lean_object* x_310; lean_object* x_311; uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; x_5 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__4; x_6 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__5; x_216 = l_Lean_Elab_Command_elabElab___closed__10; @@ -8663,13 +8663,13 @@ lean_dec(x_421); lean_dec(x_418); lean_dec_ref(x_417); lean_dec(x_416); -lean_dec(x_415); lean_dec(x_414); -lean_dec(x_412); +lean_dec(x_413); lean_dec(x_411); -lean_dec(x_410); lean_dec(x_409); -lean_dec_ref(x_406); +lean_dec(x_407); +lean_dec(x_406); +lean_dec_ref(x_405); lean_dec(x_1); x_424 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_424; @@ -8689,13 +8689,13 @@ lean_dec(x_425); lean_dec(x_418); lean_dec_ref(x_417); lean_dec(x_416); -lean_dec(x_415); lean_dec(x_414); -lean_dec(x_412); +lean_dec(x_413); lean_dec(x_411); -lean_dec(x_410); lean_dec(x_409); -lean_dec_ref(x_406); +lean_dec(x_407); +lean_dec(x_406); +lean_dec_ref(x_405); lean_dec(x_1); x_428 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_428; @@ -8703,20 +8703,20 @@ return x_428; else { lean_object* x_429; lean_object* x_430; -x_429 = l_Lean_Syntax_getArg(x_425, x_407); +x_429 = l_Lean_Syntax_getArg(x_425, x_412); lean_dec(x_425); x_430 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_430, 0, x_429); x_372 = x_405; -x_373 = x_406; -x_374 = x_409; +x_373 = x_407; +x_374 = x_406; x_375 = x_408; -x_376 = x_410; -x_377 = x_411; -x_378 = x_412; +x_376 = x_411; +x_377 = x_410; +x_378 = x_409; x_379 = x_413; -x_380 = x_416; -x_381 = x_414; +x_380 = x_414; +x_381 = x_416; x_382 = x_415; x_383 = x_430; x_384 = x_417; @@ -8732,15 +8732,15 @@ lean_object* x_431; lean_dec(x_421); x_431 = lean_box(0); x_372 = x_405; -x_373 = x_406; -x_374 = x_409; +x_373 = x_407; +x_374 = x_406; x_375 = x_408; -x_376 = x_410; -x_377 = x_411; -x_378 = x_412; +x_376 = x_411; +x_377 = x_410; +x_378 = x_409; x_379 = x_413; -x_380 = x_416; -x_381 = x_414; +x_380 = x_414; +x_381 = x_416; x_382 = x_415; x_383 = x_431; x_384 = x_417; @@ -8759,7 +8759,7 @@ if (x_449 == 0) { uint8_t x_450; lean_inc(x_448); -x_450 = l_Lean_Syntax_matchesNull(x_448, x_436); +x_450 = l_Lean_Syntax_matchesNull(x_448, x_435); if (x_450 == 0) { lean_object* x_451; @@ -8767,12 +8767,12 @@ lean_dec(x_448); lean_dec(x_445); lean_dec_ref(x_444); lean_dec(x_443); -lean_dec(x_442); lean_dec(x_441); lean_dec(x_439); lean_dec(x_438); -lean_dec(x_435); -lean_dec_ref(x_434); +lean_dec(x_436); +lean_dec(x_434); +lean_dec_ref(x_433); lean_dec(x_1); x_451 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_451; @@ -8792,12 +8792,12 @@ lean_dec(x_452); lean_dec(x_445); lean_dec_ref(x_444); lean_dec(x_443); -lean_dec(x_442); lean_dec(x_441); lean_dec(x_439); lean_dec(x_438); -lean_dec(x_435); -lean_dec_ref(x_434); +lean_dec(x_436); +lean_dec(x_434); +lean_dec_ref(x_433); lean_dec(x_1); x_455 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_455; @@ -8805,19 +8805,19 @@ return x_455; else { lean_object* x_456; lean_object* x_457; -x_456 = l_Lean_Syntax_getArg(x_452, x_437); +x_456 = l_Lean_Syntax_getArg(x_452, x_440); lean_dec(x_452); x_457 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_457, 0, x_456); x_405 = x_433; -x_406 = x_434; -x_407 = x_437; -x_408 = x_436; -x_409 = x_435; -x_410 = x_438; -x_411 = x_443; -x_412 = x_439; -x_413 = x_440; +x_406 = x_443; +x_407 = x_434; +x_408 = x_435; +x_409 = x_438; +x_410 = x_437; +x_411 = x_436; +x_412 = x_440; +x_413 = x_439; x_414 = x_441; x_415 = x_442; x_416 = x_457; @@ -8834,14 +8834,14 @@ lean_object* x_458; lean_dec(x_448); x_458 = lean_box(0); x_405 = x_433; -x_406 = x_434; -x_407 = x_437; -x_408 = x_436; -x_409 = x_435; -x_410 = x_438; -x_411 = x_443; -x_412 = x_439; -x_413 = x_440; +x_406 = x_443; +x_407 = x_434; +x_408 = x_435; +x_409 = x_438; +x_410 = x_437; +x_411 = x_436; +x_412 = x_440; +x_413 = x_439; x_414 = x_441; x_415 = x_442; x_416 = x_458; @@ -8928,16 +8928,16 @@ x_483 = l_Lean_Syntax_getArg(x_479, x_460); lean_dec(x_479); x_484 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_484, 0, x_483); -x_433 = x_474; -x_434 = x_468; -x_435 = x_461; -x_436 = x_460; -x_437 = x_472; +x_433 = x_468; +x_434 = x_467; +x_435 = x_460; +x_436 = x_469; +x_437 = x_474; x_438 = x_462; -x_439 = x_467; -x_440 = x_466; -x_441 = x_469; -x_442 = x_473; +x_439 = x_461; +x_440 = x_472; +x_441 = x_473; +x_442 = x_466; x_443 = x_484; x_444 = x_463; x_445 = x_464; @@ -8951,16 +8951,16 @@ else lean_object* x_485; lean_dec(x_475); x_485 = lean_box(0); -x_433 = x_474; -x_434 = x_468; -x_435 = x_461; -x_436 = x_460; -x_437 = x_472; +x_433 = x_468; +x_434 = x_467; +x_435 = x_460; +x_436 = x_469; +x_437 = x_474; x_438 = x_462; -x_439 = x_467; -x_440 = x_466; -x_441 = x_469; -x_442 = x_473; +x_439 = x_461; +x_440 = x_472; +x_441 = x_473; +x_442 = x_466; x_443 = x_485; x_444 = x_463; x_445 = x_464; @@ -9047,13 +9047,13 @@ goto block_486; block_47: { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_24 = l_Array_append___redArg(x_18, x_23); +x_24 = l_Array_append___redArg(x_22, x_23); lean_dec_ref(x_23); -lean_inc(x_17); +lean_inc(x_16); lean_inc(x_14); x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_17); +lean_ctor_set(x_25, 1, x_16); lean_ctor_set(x_25, 2, x_24); x_26 = l_Lean_Elab_Command_elabElabRulesAux___closed__27; lean_inc_ref(x_7); @@ -9076,83 +9076,83 @@ lean_ctor_set(x_35, 1, x_34); lean_inc(x_14); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_14); -lean_ctor_set(x_36, 1, x_11); +lean_ctor_set(x_36, 1, x_9); lean_inc(x_14); -x_37 = l_Lean_Syntax_node3(x_14, x_33, x_35, x_22, x_36); -lean_inc(x_17); +x_37 = l_Lean_Syntax_node3(x_14, x_33, x_35, x_20, x_36); +lean_inc(x_16); lean_inc(x_14); -x_38 = l_Lean_Syntax_node1(x_14, x_17, x_37); -lean_inc(x_17); +x_38 = l_Lean_Syntax_node1(x_14, x_16, x_37); +lean_inc(x_16); lean_inc(x_14); -x_39 = l_Lean_Syntax_node1(x_14, x_17, x_38); +x_39 = l_Lean_Syntax_node1(x_14, x_16, x_38); x_40 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__13; lean_inc(x_14); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_14); lean_ctor_set(x_41, 1, x_40); lean_inc(x_14); -x_42 = l_Lean_Syntax_node4(x_14, x_29, x_31, x_39, x_41, x_10); +x_42 = l_Lean_Syntax_node4(x_14, x_29, x_31, x_39, x_41, x_13); lean_inc(x_14); -x_43 = l_Lean_Syntax_node1(x_14, x_17, x_42); +x_43 = l_Lean_Syntax_node1(x_14, x_16, x_42); lean_inc(x_14); x_44 = l_Lean_Syntax_node1(x_14, x_27, x_43); -lean_inc(x_12); -x_45 = l_Lean_Syntax_node8(x_14, x_8, x_13, x_12, x_15, x_20, x_12, x_9, x_25, x_44); -x_46 = l_Lean_Elab_Command_elabCommand(x_45, x_16, x_19); +lean_inc(x_15); +x_45 = l_Lean_Syntax_node8(x_14, x_8, x_10, x_15, x_19, x_11, x_15, x_12, x_25, x_44); +x_46 = l_Lean_Elab_Command_elabCommand(x_45, x_18, x_21); return x_46; } block_77: { 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_inc_ref(x_59); -x_65 = l_Array_append___redArg(x_59, x_64); +lean_inc_ref(x_63); +x_65 = l_Array_append___redArg(x_63, x_64); lean_dec_ref(x_64); -lean_inc(x_57); lean_inc(x_56); +lean_inc(x_53); x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_56); -lean_ctor_set(x_66, 1, x_57); +lean_ctor_set(x_66, 0, x_53); +lean_ctor_set(x_66, 1, x_56); lean_ctor_set(x_66, 2, x_65); -lean_inc_ref(x_59); -lean_inc(x_57); +lean_inc_ref(x_63); lean_inc(x_56); +lean_inc(x_53); x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_56); -lean_ctor_set(x_67, 1, x_57); -lean_ctor_set(x_67, 2, x_59); +lean_ctor_set(x_67, 0, x_53); +lean_ctor_set(x_67, 1, x_56); +lean_ctor_set(x_67, 2, x_63); lean_inc_ref(x_67); -lean_inc(x_56); -x_68 = l_Lean_Syntax_node1(x_56, x_55, x_67); -lean_inc(x_56); +lean_inc(x_53); +x_68 = l_Lean_Syntax_node1(x_53, x_60, x_67); +lean_inc(x_53); x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_56); -lean_ctor_set(x_69, 1, x_52); -lean_inc(x_56); +lean_ctor_set(x_69, 0, x_53); +lean_ctor_set(x_69, 1, x_51); +lean_inc(x_53); x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_56); -lean_ctor_set(x_70, 1, x_60); -lean_inc(x_57); +lean_ctor_set(x_70, 0, x_53); +lean_ctor_set(x_70, 1, x_57); lean_inc(x_56); -x_71 = l_Lean_Syntax_node2(x_56, x_57, x_70, x_48); -if (lean_obj_tag(x_49) == 0) +lean_inc(x_53); +x_71 = l_Lean_Syntax_node2(x_53, x_56, x_70, x_54); +if (lean_obj_tag(x_55) == 0) { lean_object* x_72; x_72 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; -x_7 = x_50; -x_8 = x_51; -x_9 = x_71; -x_10 = x_53; -x_11 = x_54; -x_12 = x_67; -x_13 = x_66; -x_14 = x_56; -x_15 = x_68; -x_16 = x_58; -x_17 = x_57; +x_7 = x_48; +x_8 = x_49; +x_9 = x_50; +x_10 = x_66; +x_11 = x_69; +x_12 = x_71; +x_13 = x_52; +x_14 = x_53; +x_15 = x_67; +x_16 = x_56; +x_17 = lean_box(0); x_18 = x_59; -x_19 = x_61; -x_20 = x_69; -x_21 = lean_box(0); +x_19 = x_68; +x_20 = x_61; +x_21 = x_62; x_22 = x_63; x_23 = x_72; goto block_47; @@ -9160,30 +9160,30 @@ goto block_47; else { lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_73 = lean_ctor_get(x_49, 0); +x_73 = lean_ctor_get(x_55, 0); lean_inc(x_73); -lean_dec_ref(x_49); +lean_dec_ref(x_55); x_74 = l_Lean_Elab_Command_elabElabRules___lam__0___closed__0; -lean_inc(x_56); +lean_inc(x_53); x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_56); +lean_ctor_set(x_75, 0, x_53); lean_ctor_set(x_75, 1, x_74); x_76 = l_Array_mkArray2___redArg(x_75, x_73); -x_7 = x_50; -x_8 = x_51; -x_9 = x_71; -x_10 = x_53; -x_11 = x_54; -x_12 = x_67; -x_13 = x_66; -x_14 = x_56; -x_15 = x_68; -x_16 = x_58; -x_17 = x_57; +x_7 = x_48; +x_8 = x_49; +x_9 = x_50; +x_10 = x_66; +x_11 = x_69; +x_12 = x_71; +x_13 = x_52; +x_14 = x_53; +x_15 = x_67; +x_16 = x_56; +x_17 = lean_box(0); x_18 = x_59; -x_19 = x_61; -x_20 = x_69; -x_21 = lean_box(0); +x_19 = x_68; +x_20 = x_61; +x_21 = x_62; x_22 = x_63; x_23 = x_76; goto block_47; @@ -9194,25 +9194,25 @@ goto block_47; lean_object* x_93; lean_object* x_94; x_93 = l_Lean_Elab_Command_elabElabRules___lam__1___closed__0; x_94 = l_Lean_Elab_Command_elabElabRules___lam__1___closed__1; -if (lean_obj_tag(x_81) == 0) +if (lean_obj_tag(x_88) == 0) { lean_object* x_95; x_95 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; x_48 = x_78; -x_49 = x_79; -x_50 = x_80; -x_51 = x_94; -x_52 = x_93; -x_53 = x_82; -x_54 = x_83; -x_55 = x_84; -x_56 = x_85; -x_57 = x_86; -x_58 = x_87; -x_59 = x_88; -x_60 = x_89; -x_61 = x_90; -x_62 = lean_box(0); +x_49 = x_94; +x_50 = x_79; +x_51 = x_93; +x_52 = x_80; +x_53 = x_81; +x_54 = x_82; +x_55 = x_83; +x_56 = x_84; +x_57 = x_85; +x_58 = lean_box(0); +x_59 = x_86; +x_60 = x_87; +x_61 = x_89; +x_62 = x_90; x_63 = x_91; x_64 = x_95; goto block_77; @@ -9220,25 +9220,25 @@ goto block_77; else { lean_object* x_96; lean_object* x_97; -x_96 = lean_ctor_get(x_81, 0); +x_96 = lean_ctor_get(x_88, 0); lean_inc(x_96); -lean_dec_ref(x_81); +lean_dec_ref(x_88); x_97 = l_Array_mkArray1___redArg(x_96); x_48 = x_78; -x_49 = x_79; -x_50 = x_80; -x_51 = x_94; -x_52 = x_93; -x_53 = x_82; -x_54 = x_83; -x_55 = x_84; -x_56 = x_85; -x_57 = x_86; -x_58 = x_87; -x_59 = x_88; -x_60 = x_89; -x_61 = x_90; -x_62 = lean_box(0); +x_49 = x_94; +x_50 = x_79; +x_51 = x_93; +x_52 = x_80; +x_53 = x_81; +x_54 = x_82; +x_55 = x_83; +x_56 = x_84; +x_57 = x_85; +x_58 = lean_box(0); +x_59 = x_86; +x_60 = x_87; +x_61 = x_89; +x_62 = x_90; x_63 = x_91; x_64 = x_97; goto block_77; @@ -9247,142 +9247,142 @@ goto block_77; block_175: { 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; size_t 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_inc_ref(x_117); -x_123 = l_Array_append___redArg(x_117, x_122); +lean_inc_ref(x_121); +x_123 = l_Array_append___redArg(x_121, x_122); lean_dec_ref(x_122); -lean_inc(x_115); -lean_inc(x_110); +lean_inc(x_108); +lean_inc(x_109); x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_110); -lean_ctor_set(x_124, 1, x_115); +lean_ctor_set(x_124, 0, x_109); +lean_ctor_set(x_124, 1, x_108); lean_ctor_set(x_124, 2, x_123); x_125 = l_Lean_Elab_Command_elabElab___closed__3; x_126 = l_Lean_Elab_Command_elabElabRules___lam__0___closed__1; -lean_inc(x_110); +lean_inc(x_109); x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_110); +lean_ctor_set(x_127, 0, x_109); lean_ctor_set(x_127, 1, x_126); x_128 = l_Lean_Elab_Command_elabElab___closed__4; -lean_inc(x_110); +lean_inc(x_109); x_129 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_129, 0, x_110); +lean_ctor_set(x_129, 0, x_109); lean_ctor_set(x_129, 1, x_128); x_130 = l_Lean_Elab_Command_elabElabRulesAux___closed__11; -lean_inc(x_110); +lean_inc(x_109); x_131 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_131, 0, x_110); +lean_ctor_set(x_131, 0, x_109); lean_ctor_set(x_131, 1, x_130); -x_132 = l_Nat_reprFast(x_104); +x_132 = l_Nat_reprFast(x_101); x_133 = lean_box(2); x_134 = l_Lean_Syntax_mkNumLit(x_132, x_133); x_135 = l_Lean_Elab_Command_elabElabRules___lam__0___closed__3; -lean_inc(x_110); +lean_inc(x_109); x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_110); +lean_ctor_set(x_136, 0, x_109); lean_ctor_set(x_136, 1, x_135); -lean_inc(x_110); -x_137 = l_Lean_Syntax_node5(x_110, x_125, x_127, x_129, x_131, x_134, x_136); -lean_inc(x_115); -lean_inc(x_110); -x_138 = l_Lean_Syntax_node1(x_110, x_115, x_137); -x_139 = lean_array_size(x_103); -x_140 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElab_spec__13(x_139, x_112, x_103); -lean_inc_ref(x_117); -x_141 = l_Array_append___redArg(x_117, x_140); +lean_inc(x_109); +x_137 = l_Lean_Syntax_node5(x_109, x_125, x_127, x_129, x_131, x_134, x_136); +lean_inc(x_108); +lean_inc(x_109); +x_138 = l_Lean_Syntax_node1(x_109, x_108, x_137); +x_139 = lean_array_size(x_115); +x_140 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElab_spec__13(x_139, x_117, x_115); +lean_inc_ref(x_121); +x_141 = l_Array_append___redArg(x_121, x_140); lean_dec_ref(x_140); -lean_inc(x_115); -lean_inc(x_110); +lean_inc(x_108); +lean_inc(x_109); x_142 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_142, 0, x_110); -lean_ctor_set(x_142, 1, x_115); +lean_ctor_set(x_142, 0, x_109); +lean_ctor_set(x_142, 1, x_108); lean_ctor_set(x_142, 2, x_141); x_143 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_110); +lean_inc(x_109); x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_110); +lean_ctor_set(x_144, 0, x_109); lean_ctor_set(x_144, 1, x_143); x_145 = l_Lean_Elab_Command_elabElab___closed__5; -x_146 = lean_array_push(x_145, x_116); -x_147 = lean_array_push(x_146, x_111); -x_148 = lean_array_push(x_147, x_120); -x_149 = lean_array_push(x_148, x_108); -x_150 = lean_array_push(x_149, x_107); +x_146 = lean_array_push(x_145, x_105); +x_147 = lean_array_push(x_146, x_116); +x_148 = lean_array_push(x_147, x_111); +x_149 = lean_array_push(x_148, x_104); +x_150 = lean_array_push(x_149, x_120); x_151 = lean_array_push(x_150, x_124); x_152 = lean_array_push(x_151, x_138); x_153 = lean_array_push(x_152, x_142); x_154 = lean_array_push(x_153, x_144); -lean_inc(x_99); -x_155 = lean_array_push(x_154, x_99); +lean_inc(x_106); +x_155 = lean_array_push(x_154, x_106); x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_110); -lean_ctor_set(x_156, 1, x_121); +lean_ctor_set(x_156, 0, x_109); +lean_ctor_set(x_156, 1, x_112); lean_ctor_set(x_156, 2, x_155); -lean_inc(x_118); -lean_inc_ref(x_114); -x_157 = l_Lean_Elab_Command_elabSyntax(x_156, x_114, x_118); +lean_inc(x_119); +lean_inc_ref(x_110); +x_157 = l_Lean_Elab_Command_elabSyntax(x_156, x_110, x_119); if (lean_obj_tag(x_157) == 0) { lean_object* x_158; lean_object* x_159; x_158 = lean_ctor_get(x_157, 0); lean_inc(x_158); lean_dec_ref(x_157); -x_159 = l_Lean_Elab_Command_getRef___redArg(x_114); +x_159 = l_Lean_Elab_Command_getRef___redArg(x_110); if (lean_obj_tag(x_159) == 0) { lean_object* x_160; lean_object* x_161; x_160 = lean_ctor_get(x_159, 0); lean_inc(x_160); lean_dec_ref(x_159); -x_161 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_114); +x_161 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_110); if (lean_obj_tag(x_161) == 0) { lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_dec_ref(x_161); -x_162 = lean_ctor_get(x_114, 5); +x_162 = lean_ctor_get(x_110, 5); x_163 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_163, 0, x_133); lean_ctor_set(x_163, 1, x_158); -lean_ctor_set(x_163, 2, x_100); -x_164 = l_Lean_SourceInfo_fromRef(x_160, x_106); +lean_ctor_set(x_163, 2, x_102); +x_164 = l_Lean_SourceInfo_fromRef(x_160, x_118); lean_dec(x_160); if (lean_obj_tag(x_162) == 0) { lean_object* x_165; -x_165 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_118); +x_165 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_119); lean_dec_ref(x_165); x_78 = x_99; -x_79 = x_102; -x_80 = x_101; -x_81 = x_105; -x_82 = x_109; -x_83 = x_135; -x_84 = x_113; -x_85 = x_164; -x_86 = x_115; -x_87 = x_114; -x_88 = x_117; -x_89 = x_143; -x_90 = x_118; -x_91 = x_163; +x_79 = x_135; +x_80 = x_103; +x_81 = x_164; +x_82 = x_106; +x_83 = x_107; +x_84 = x_108; +x_85 = x_143; +x_86 = x_110; +x_87 = x_113; +x_88 = x_114; +x_89 = x_163; +x_90 = x_119; +x_91 = x_121; x_92 = lean_box(0); goto block_98; } else { x_78 = x_99; -x_79 = x_102; -x_80 = x_101; -x_81 = x_105; -x_82 = x_109; -x_83 = x_135; -x_84 = x_113; -x_85 = x_164; -x_86 = x_115; -x_87 = x_114; -x_88 = x_117; -x_89 = x_143; -x_90 = x_118; -x_91 = x_163; +x_79 = x_135; +x_80 = x_103; +x_81 = x_164; +x_82 = x_106; +x_83 = x_107; +x_84 = x_108; +x_85 = x_143; +x_86 = x_110; +x_87 = x_113; +x_88 = x_114; +x_89 = x_163; +x_90 = x_119; +x_91 = x_121; x_92 = lean_box(0); goto block_98; } @@ -9392,17 +9392,17 @@ else uint8_t x_166; lean_dec(x_160); lean_dec(x_158); -lean_dec(x_118); -lean_dec_ref(x_117); -lean_dec(x_115); -lean_dec_ref(x_114); +lean_dec_ref(x_121); +lean_dec(x_119); +lean_dec(x_114); lean_dec(x_113); -lean_dec(x_109); -lean_dec(x_105); -lean_dec(x_102); -lean_dec_ref(x_101); -lean_dec_ref(x_100); -lean_dec(x_99); +lean_dec_ref(x_110); +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_106); +lean_dec(x_103); +lean_dec_ref(x_102); +lean_dec_ref(x_99); x_166 = !lean_is_exclusive(x_161); if (x_166 == 0) { @@ -9424,17 +9424,17 @@ else { uint8_t x_169; lean_dec(x_158); -lean_dec(x_118); -lean_dec_ref(x_117); -lean_dec(x_115); -lean_dec_ref(x_114); +lean_dec_ref(x_121); +lean_dec(x_119); +lean_dec(x_114); lean_dec(x_113); -lean_dec(x_109); -lean_dec(x_105); -lean_dec(x_102); -lean_dec_ref(x_101); -lean_dec_ref(x_100); -lean_dec(x_99); +lean_dec_ref(x_110); +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_106); +lean_dec(x_103); +lean_dec_ref(x_102); +lean_dec_ref(x_99); x_169 = !lean_is_exclusive(x_159); if (x_169 == 0) { @@ -9455,17 +9455,17 @@ return x_171; else { uint8_t x_172; -lean_dec(x_118); -lean_dec_ref(x_117); -lean_dec(x_115); -lean_dec_ref(x_114); +lean_dec_ref(x_121); +lean_dec(x_119); +lean_dec(x_114); lean_dec(x_113); -lean_dec(x_109); -lean_dec(x_105); -lean_dec(x_102); -lean_dec_ref(x_101); -lean_dec_ref(x_100); -lean_dec(x_99); +lean_dec_ref(x_110); +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_106); +lean_dec(x_103); +lean_dec_ref(x_102); +lean_dec_ref(x_99); x_172 = !lean_is_exclusive(x_157); if (x_172 == 0) { @@ -9486,41 +9486,41 @@ return x_174; block_215: { lean_object* x_200; lean_object* x_201; -lean_inc_ref(x_194); -x_200 = l_Array_append___redArg(x_194, x_199); +lean_inc_ref(x_198); +x_200 = l_Array_append___redArg(x_198, x_199); lean_dec_ref(x_199); -lean_inc(x_193); +lean_inc(x_185); lean_inc(x_186); x_201 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_201, 0, x_186); -lean_ctor_set(x_201, 1, x_193); +lean_ctor_set(x_201, 1, x_185); lean_ctor_set(x_201, 2, x_200); -if (lean_obj_tag(x_189) == 0) +if (lean_obj_tag(x_194) == 0) { lean_object* x_202; x_202 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; x_99 = x_176; -x_100 = x_177; +x_100 = lean_box(0); x_101 = x_178; x_102 = x_179; x_103 = x_180; x_104 = x_181; x_105 = x_182; x_106 = x_183; -x_107 = x_201; -x_108 = x_184; -x_109 = x_185; -x_110 = x_186; +x_107 = x_184; +x_108 = x_185; +x_109 = x_186; +x_110 = x_188; x_111 = x_187; -x_112 = x_188; +x_112 = x_189; x_113 = x_190; -x_114 = x_192; +x_114 = x_191; x_115 = x_193; -x_116 = x_191; -x_117 = x_194; -x_118 = x_195; -x_119 = lean_box(0); -x_120 = x_197; +x_116 = x_192; +x_117 = x_195; +x_118 = x_197; +x_119 = x_196; +x_120 = x_201; x_121 = x_198; x_122 = x_202; goto block_175; @@ -9528,9 +9528,9 @@ goto block_175; else { lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -x_203 = lean_ctor_get(x_189, 0); +x_203 = lean_ctor_get(x_194, 0); lean_inc(x_203); -lean_dec_ref(x_189); +lean_dec_ref(x_194); x_204 = l_Lean_Elab_Command_elabElab___closed__7; x_205 = l_Lean_Elab_Command_elabElabRules___lam__0___closed__1; lean_inc(x_186); @@ -9556,27 +9556,27 @@ lean_inc(x_186); x_213 = l_Lean_Syntax_node5(x_186, x_204, x_206, x_208, x_210, x_203, x_212); x_214 = l_Array_mkArray1___redArg(x_213); x_99 = x_176; -x_100 = x_177; +x_100 = lean_box(0); x_101 = x_178; x_102 = x_179; x_103 = x_180; x_104 = x_181; x_105 = x_182; x_106 = x_183; -x_107 = x_201; -x_108 = x_184; -x_109 = x_185; -x_110 = x_186; +x_107 = x_184; +x_108 = x_185; +x_109 = x_186; +x_110 = x_188; x_111 = x_187; -x_112 = x_188; +x_112 = x_189; x_113 = x_190; -x_114 = x_192; +x_114 = x_191; x_115 = x_193; -x_116 = x_191; -x_117 = x_194; -x_118 = x_195; -x_119 = lean_box(0); -x_120 = x_197; +x_116 = x_192; +x_117 = x_195; +x_118 = x_197; +x_119 = x_196; +x_120 = x_201; x_121 = x_198; x_122 = x_214; goto block_175; @@ -9585,88 +9585,88 @@ goto block_175; block_254: { lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_inc_ref(x_235); -x_243 = l_Array_append___redArg(x_235, x_242); +lean_inc_ref(x_241); +x_243 = l_Array_append___redArg(x_241, x_242); lean_dec_ref(x_242); -lean_inc(x_234); lean_inc(x_228); +lean_inc(x_229); x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_228); -lean_ctor_set(x_244, 1, x_234); +lean_ctor_set(x_244, 0, x_229); +lean_ctor_set(x_244, 1, x_228); lean_ctor_set(x_244, 2, x_243); -x_245 = l_Lean_SourceInfo_fromRef(x_241, x_217); -lean_dec(x_241); +x_245 = l_Lean_SourceInfo_fromRef(x_238, x_217); +lean_dec(x_238); x_246 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_246, 0, x_245); -lean_ctor_set(x_246, 1, x_224); -if (lean_obj_tag(x_238) == 0) +lean_ctor_set(x_246, 1, x_227); +if (lean_obj_tag(x_221) == 0) { lean_object* x_247; x_247 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; x_176 = x_218; -x_177 = x_219; +x_177 = lean_box(0); x_178 = x_220; -x_179 = x_221; -x_180 = x_222; -x_181 = x_223; -x_182 = x_225; -x_183 = x_226; -x_184 = x_246; -x_185 = x_227; -x_186 = x_228; -x_187 = x_244; -x_188 = x_229; -x_189 = x_231; -x_190 = x_230; -x_191 = x_233; -x_192 = x_232; -x_193 = x_234; -x_194 = x_235; +x_179 = x_222; +x_180 = x_223; +x_181 = x_246; +x_182 = x_224; +x_183 = x_225; +x_184 = x_226; +x_185 = x_228; +x_186 = x_229; +x_187 = x_231; +x_188 = x_230; +x_189 = x_232; +x_190 = x_233; +x_191 = x_234; +x_192 = x_244; +x_193 = x_235; +x_194 = x_237; x_195 = x_236; -x_196 = lean_box(0); +x_196 = x_240; x_197 = x_239; -x_198 = x_240; +x_198 = x_241; x_199 = x_247; goto block_215; } 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; -x_248 = lean_ctor_get(x_238, 0); +x_248 = lean_ctor_get(x_221, 0); lean_inc(x_248); -lean_dec_ref(x_238); +lean_dec_ref(x_221); x_249 = l_Lean_Elab_Command_elabElab___closed__12; x_250 = l_Lean_Elab_Command_elabElabRulesAux___closed__7; -lean_inc(x_228); +lean_inc(x_229); x_251 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_251, 0, x_228); +lean_ctor_set(x_251, 0, x_229); lean_ctor_set(x_251, 1, x_250); -lean_inc(x_228); -x_252 = l_Lean_Syntax_node2(x_228, x_249, x_251, x_248); +lean_inc(x_229); +x_252 = l_Lean_Syntax_node2(x_229, x_249, x_251, x_248); x_253 = l_Array_mkArray1___redArg(x_252); x_176 = x_218; -x_177 = x_219; +x_177 = lean_box(0); x_178 = x_220; -x_179 = x_221; -x_180 = x_222; -x_181 = x_223; -x_182 = x_225; -x_183 = x_226; -x_184 = x_246; -x_185 = x_227; -x_186 = x_228; -x_187 = x_244; -x_188 = x_229; -x_189 = x_231; -x_190 = x_230; -x_191 = x_233; -x_192 = x_232; -x_193 = x_234; -x_194 = x_235; +x_179 = x_222; +x_180 = x_223; +x_181 = x_246; +x_182 = x_224; +x_183 = x_225; +x_184 = x_226; +x_185 = x_228; +x_186 = x_229; +x_187 = x_231; +x_188 = x_230; +x_189 = x_232; +x_190 = x_233; +x_191 = x_234; +x_192 = x_244; +x_193 = x_235; +x_194 = x_237; x_195 = x_236; -x_196 = lean_box(0); +x_196 = x_240; x_197 = x_239; -x_198 = x_240; +x_198 = x_241; x_199 = x_253; goto block_215; } @@ -9674,42 +9674,42 @@ goto block_215; block_294: { lean_object* x_280; lean_object* x_281; -lean_inc_ref(x_272); -x_280 = l_Array_append___redArg(x_272, x_279); +lean_inc_ref(x_278); +x_280 = l_Array_append___redArg(x_278, x_279); lean_dec_ref(x_279); -lean_inc(x_271); +lean_inc(x_265); lean_inc(x_266); x_281 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_281, 0, x_266); -lean_ctor_set(x_281, 1, x_271); +lean_ctor_set(x_281, 1, x_265); lean_ctor_set(x_281, 2, x_280); -if (lean_obj_tag(x_264) == 0) +if (lean_obj_tag(x_259) == 0) { lean_object* x_282; x_282 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; x_218 = x_255; -x_219 = x_256; +x_219 = lean_box(0); x_220 = x_257; x_221 = x_258; -x_222 = x_259; -x_223 = x_260; -x_224 = x_261; +x_222 = x_260; +x_223 = x_261; +x_224 = x_281; x_225 = x_262; x_226 = x_263; -x_227 = x_265; -x_228 = x_266; -x_229 = x_267; -x_230 = x_269; -x_231 = x_268; -x_232 = x_270; -x_233 = x_281; +x_227 = x_264; +x_228 = x_265; +x_229 = x_266; +x_230 = x_268; +x_231 = x_267; +x_232 = x_269; +x_233 = x_270; x_234 = x_271; x_235 = x_272; -x_236 = x_273; -x_237 = lean_box(0); -x_238 = x_275; -x_239 = x_276; -x_240 = x_277; +x_236 = x_275; +x_237 = x_274; +x_238 = x_273; +x_239 = x_277; +x_240 = x_276; x_241 = x_278; x_242 = x_282; goto block_254; @@ -9717,25 +9717,25 @@ goto block_254; else { lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_283 = lean_ctor_get(x_264, 0); +x_283 = lean_ctor_get(x_259, 0); lean_inc(x_283); -lean_dec_ref(x_264); +lean_dec_ref(x_259); x_284 = l_Lean_Elab_Command_elabElabRulesAux___closed__0; -lean_inc_ref(x_257); -x_285 = l_Lean_Name_mkStr4(x_5, x_6, x_257, x_284); +lean_inc_ref(x_255); +x_285 = l_Lean_Name_mkStr4(x_5, x_6, x_255, x_284); x_286 = l_Lean_Elab_Command_elabElabRulesAux___closed__1; lean_inc(x_266); x_287 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_287, 0, x_266); lean_ctor_set(x_287, 1, x_286); -lean_inc_ref(x_272); -x_288 = l_Array_append___redArg(x_272, x_283); +lean_inc_ref(x_278); +x_288 = l_Array_append___redArg(x_278, x_283); lean_dec(x_283); -lean_inc(x_271); +lean_inc(x_265); lean_inc(x_266); x_289 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_289, 0, x_266); -lean_ctor_set(x_289, 1, x_271); +lean_ctor_set(x_289, 1, x_265); lean_ctor_set(x_289, 2, x_288); x_290 = l_Lean_Elab_Command_elabElabRulesAux___closed__3; lean_inc(x_266); @@ -9746,28 +9746,28 @@ lean_inc(x_266); x_292 = l_Lean_Syntax_node3(x_266, x_285, x_287, x_289, x_291); x_293 = l_Array_mkArray1___redArg(x_292); x_218 = x_255; -x_219 = x_256; +x_219 = lean_box(0); x_220 = x_257; x_221 = x_258; -x_222 = x_259; -x_223 = x_260; -x_224 = x_261; +x_222 = x_260; +x_223 = x_261; +x_224 = x_281; x_225 = x_262; x_226 = x_263; -x_227 = x_265; -x_228 = x_266; -x_229 = x_267; -x_230 = x_269; -x_231 = x_268; -x_232 = x_270; -x_233 = x_281; +x_227 = x_264; +x_228 = x_265; +x_229 = x_266; +x_230 = x_268; +x_231 = x_267; +x_232 = x_269; +x_233 = x_270; x_234 = x_271; x_235 = x_272; -x_236 = x_273; -x_237 = lean_box(0); -x_238 = x_275; -x_239 = x_276; -x_240 = x_277; +x_236 = x_275; +x_237 = x_274; +x_238 = x_273; +x_239 = x_277; +x_240 = x_276; x_241 = x_278; x_242 = x_293; goto block_254; @@ -9780,67 +9780,67 @@ x_315 = l_Lean_Elab_Command_elabElab___closed__13; x_316 = l_Lean_Elab_Command_elabElab___closed__14; x_317 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__11; x_318 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabElabRulesAux_spec__10_spec__10___closed__12; -if (lean_obj_tag(x_301) == 0) +if (lean_obj_tag(x_307) == 0) { lean_object* x_319; x_319 = l_Lean_Elab_Command_elabElabRulesAux___closed__40; x_255 = x_295; -x_256 = x_296; -x_257 = x_297; -x_258 = x_298; -x_259 = x_299; -x_260 = x_300; -x_261 = x_315; +x_256 = lean_box(0); +x_257 = x_296; +x_258 = x_297; +x_259 = x_298; +x_260 = x_299; +x_261 = x_300; x_262 = x_301; x_263 = x_302; -x_264 = x_303; -x_265 = x_304; -x_266 = x_305; -x_267 = x_306; -x_268 = x_308; -x_269 = x_307; -x_270 = x_309; -x_271 = x_317; -x_272 = x_318; -x_273 = x_310; -x_274 = lean_box(0); -x_275 = x_311; -x_276 = x_312; -x_277 = x_316; -x_278 = x_313; +x_264 = x_315; +x_265 = x_317; +x_266 = x_303; +x_267 = x_304; +x_268 = x_305; +x_269 = x_316; +x_270 = x_306; +x_271 = x_307; +x_272 = x_308; +x_273 = x_311; +x_274 = x_310; +x_275 = x_309; +x_276 = x_313; +x_277 = x_312; +x_278 = x_318; x_279 = x_319; goto block_294; } else { lean_object* x_320; lean_object* x_321; -x_320 = lean_ctor_get(x_301, 0); +x_320 = lean_ctor_get(x_307, 0); lean_inc(x_320); x_321 = l_Array_mkArray1___redArg(x_320); x_255 = x_295; -x_256 = x_296; -x_257 = x_297; -x_258 = x_298; -x_259 = x_299; -x_260 = x_300; -x_261 = x_315; +x_256 = lean_box(0); +x_257 = x_296; +x_258 = x_297; +x_259 = x_298; +x_260 = x_299; +x_261 = x_300; x_262 = x_301; x_263 = x_302; -x_264 = x_303; -x_265 = x_304; -x_266 = x_305; -x_267 = x_306; -x_268 = x_308; -x_269 = x_307; -x_270 = x_309; -x_271 = x_317; -x_272 = x_318; -x_273 = x_310; -x_274 = lean_box(0); -x_275 = x_311; -x_276 = x_312; -x_277 = x_316; -x_278 = x_313; +x_264 = x_315; +x_265 = x_317; +x_266 = x_303; +x_267 = x_304; +x_268 = x_305; +x_269 = x_316; +x_270 = x_306; +x_271 = x_307; +x_272 = x_308; +x_273 = x_311; +x_274 = x_310; +x_275 = x_309; +x_276 = x_313; +x_277 = x_312; +x_278 = x_318; x_279 = x_321; goto block_294; } @@ -9849,7 +9849,7 @@ goto block_294; { lean_object* x_340; lean_object* x_341; x_340 = lean_alloc_closure((void*)(l_Lean_evalOptPrio), 3, 1); -lean_closure_set(x_340, 0, x_329); +lean_closure_set(x_340, 0, x_325); lean_inc_ref(x_337); x_341 = l_Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabElab_spec__0___redArg(x_340, x_337, x_338); if (lean_obj_tag(x_341) == 0) @@ -9858,8 +9858,8 @@ lean_object* x_342; lean_object* x_343; size_t x_344; size_t x_345; lean_object* x_342 = lean_ctor_get(x_341, 0); lean_inc(x_342); lean_dec_ref(x_341); -x_343 = l_Lean_Syntax_getArgs(x_327); -lean_dec(x_327); +x_343 = l_Lean_Syntax_getArgs(x_326); +lean_dec(x_326); x_344 = lean_array_size(x_343); x_345 = 0; lean_inc_ref(x_337); @@ -9890,8 +9890,8 @@ if (lean_obj_tag(x_353) == 0) lean_object* x_354; lean_object* x_355; uint8_t x_356; lean_object* x_357; lean_dec_ref(x_353); x_354 = lean_ctor_get(x_337, 5); -x_355 = l_Lean_Syntax_getArg(x_332, x_323); -lean_dec(x_332); +x_355 = l_Lean_Syntax_getArg(x_333, x_330); +lean_dec(x_333); x_356 = 0; x_357 = l_Lean_SourceInfo_fromRef(x_352, x_356); lean_dec(x_352); @@ -9900,49 +9900,49 @@ if (lean_obj_tag(x_354) == 0) lean_object* x_358; x_358 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabElabRulesAux_spec__1___redArg(x_338); lean_dec_ref(x_358); -x_295 = x_324; -x_296 = x_350; -x_297 = x_325; -x_298 = x_336; -x_299 = x_349; -x_300 = x_342; -x_301 = x_326; -x_302 = x_356; -x_303 = x_328; -x_304 = x_355; -x_305 = x_357; -x_306 = x_345; -x_307 = x_330; -x_308 = x_331; -x_309 = x_337; -x_310 = x_338; -x_311 = x_333; -x_312 = x_334; -x_313 = x_335; +x_295 = x_323; +x_296 = x_342; +x_297 = x_324; +x_298 = x_327; +x_299 = x_350; +x_300 = x_355; +x_301 = x_328; +x_302 = x_336; +x_303 = x_357; +x_304 = x_329; +x_305 = x_337; +x_306 = x_331; +x_307 = x_332; +x_308 = x_349; +x_309 = x_345; +x_310 = x_334; +x_311 = x_335; +x_312 = x_356; +x_313 = x_338; x_314 = lean_box(0); goto block_322; } else { -x_295 = x_324; -x_296 = x_350; -x_297 = x_325; -x_298 = x_336; -x_299 = x_349; -x_300 = x_342; -x_301 = x_326; -x_302 = x_356; -x_303 = x_328; -x_304 = x_355; -x_305 = x_357; -x_306 = x_345; -x_307 = x_330; -x_308 = x_331; -x_309 = x_337; -x_310 = x_338; -x_311 = x_333; -x_312 = x_334; -x_313 = x_335; +x_295 = x_323; +x_296 = x_342; +x_297 = x_324; +x_298 = x_327; +x_299 = x_350; +x_300 = x_355; +x_301 = x_328; +x_302 = x_336; +x_303 = x_357; +x_304 = x_329; +x_305 = x_337; +x_306 = x_331; +x_307 = x_332; +x_308 = x_349; +x_309 = x_345; +x_310 = x_334; +x_311 = x_335; +x_312 = x_356; +x_313 = x_338; x_314 = lean_box(0); goto block_322; } @@ -9962,11 +9962,11 @@ lean_dec(x_334); lean_dec(x_333); lean_dec(x_332); lean_dec(x_331); -lean_dec(x_330); +lean_dec(x_329); lean_dec(x_328); -lean_dec(x_326); -lean_dec_ref(x_325); +lean_dec(x_327); lean_dec(x_324); +lean_dec_ref(x_323); x_359 = !lean_is_exclusive(x_353); if (x_359 == 0) { @@ -9998,11 +9998,11 @@ lean_dec(x_334); lean_dec(x_333); lean_dec(x_332); lean_dec(x_331); -lean_dec(x_330); +lean_dec(x_329); lean_dec(x_328); -lean_dec(x_326); -lean_dec_ref(x_325); +lean_dec(x_327); lean_dec(x_324); +lean_dec_ref(x_323); x_362 = !lean_is_exclusive(x_351); if (x_362 == 0) { @@ -10032,11 +10032,11 @@ lean_dec(x_334); lean_dec(x_333); lean_dec(x_332); lean_dec(x_331); -lean_dec(x_330); +lean_dec(x_329); lean_dec(x_328); -lean_dec(x_326); -lean_dec_ref(x_325); +lean_dec(x_327); lean_dec(x_324); +lean_dec_ref(x_323); x_365 = !lean_is_exclusive(x_346); if (x_365 == 0) { @@ -10065,12 +10065,12 @@ lean_dec(x_334); lean_dec(x_333); lean_dec(x_332); lean_dec(x_331); -lean_dec(x_330); +lean_dec(x_329); lean_dec(x_328); lean_dec(x_327); lean_dec(x_326); -lean_dec_ref(x_325); lean_dec(x_324); +lean_dec_ref(x_323); x_368 = !lean_is_exclusive(x_341); if (x_368 == 0) { @@ -10103,14 +10103,14 @@ lean_dec(x_388); lean_dec(x_385); lean_dec_ref(x_384); lean_dec(x_383); -lean_dec(x_382); lean_dec(x_381); lean_dec(x_380); +lean_dec(x_379); lean_dec(x_378); -lean_dec(x_377); lean_dec(x_376); lean_dec(x_374); -lean_dec_ref(x_373); +lean_dec(x_373); +lean_dec_ref(x_372); lean_dec(x_1); x_391 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_391; @@ -10122,13 +10122,13 @@ x_392 = lean_unsigned_to_nat(7u); x_393 = l_Lean_Syntax_getArg(x_1, x_392); lean_dec(x_1); x_394 = l_Lean_Syntax_getArg(x_388, x_375); -x_395 = l_Lean_Syntax_getArg(x_388, x_379); +x_395 = l_Lean_Syntax_getArg(x_388, x_382); x_396 = l_Lean_Syntax_isNone(x_395); if (x_396 == 0) { uint8_t x_397; lean_inc(x_395); -x_397 = l_Lean_Syntax_matchesNull(x_395, x_379); +x_397 = l_Lean_Syntax_matchesNull(x_395, x_382); if (x_397 == 0) { lean_object* x_398; @@ -10139,14 +10139,14 @@ lean_dec(x_388); lean_dec(x_385); lean_dec_ref(x_384); lean_dec(x_383); -lean_dec(x_382); lean_dec(x_381); lean_dec(x_380); +lean_dec(x_379); lean_dec(x_378); -lean_dec(x_377); lean_dec(x_376); lean_dec(x_374); -lean_dec_ref(x_373); +lean_dec(x_373); +lean_dec_ref(x_372); x_398 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabElabRulesAux_spec__2___redArg(); return x_398; } @@ -10158,18 +10158,18 @@ lean_dec(x_395); x_400 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_400, 0, x_399); x_323 = x_372; -x_324 = x_394; -x_325 = x_373; -x_326 = x_374; -x_327 = x_393; -x_328 = x_376; -x_329 = x_383; -x_330 = x_381; -x_331 = x_380; -x_332 = x_388; -x_333 = x_377; -x_334 = x_378; -x_335 = x_382; +x_324 = x_374; +x_325 = x_383; +x_326 = x_393; +x_327 = x_378; +x_328 = x_394; +x_329 = x_373; +x_330 = x_377; +x_331 = x_376; +x_332 = x_379; +x_333 = x_388; +x_334 = x_381; +x_335 = x_380; x_336 = x_400; x_337 = x_384; x_338 = x_385; @@ -10183,18 +10183,18 @@ lean_object* x_401; lean_dec(x_395); x_401 = lean_box(0); x_323 = x_372; -x_324 = x_394; -x_325 = x_373; -x_326 = x_374; -x_327 = x_393; -x_328 = x_376; -x_329 = x_383; -x_330 = x_381; -x_331 = x_380; -x_332 = x_388; -x_333 = x_377; -x_334 = x_378; -x_335 = x_382; +x_324 = x_374; +x_325 = x_383; +x_326 = x_393; +x_327 = x_378; +x_328 = x_394; +x_329 = x_373; +x_330 = x_377; +x_331 = x_376; +x_332 = x_379; +x_333 = x_388; +x_334 = x_381; +x_335 = x_380; x_336 = x_401; x_337 = x_384; x_338 = x_385; diff --git a/stage0/stdlib/Lean/Elab/ErrorExplanation.c b/stage0/stdlib/Lean/Elab/ErrorExplanation.c index 7ec3d8ac6582..acf02bb8816a 100644 --- a/stage0/stdlib/Lean/Elab/ErrorExplanation.c +++ b/stage0/stdlib/Lean/Elab/ErrorExplanation.c @@ -5383,7 +5383,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22_spec__22___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; 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; uint8_t x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_77; uint8_t x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_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; lean_object* x_11; uint8_t 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; lean_object* x_50; uint8_t x_51; lean_object* 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; 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; lean_object* x_89; uint8_t x_90; lean_object* 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; 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) @@ -5418,15 +5418,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_16); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_16); -lean_ctor_set(x_27, 1, x_10); -lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_12); +lean_ctor_set(x_27, 0, x_13); +lean_ctor_set(x_27, 1, x_15); +lean_ctor_set(x_27, 2, x_11); +lean_ctor_set(x_27, 3, x_10); lean_ctor_set(x_27, 4, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_13); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_11); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_14); 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); @@ -5463,15 +5463,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_16); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_16); -lean_ctor_set(x_43, 1, x_10); -lean_ctor_set(x_43, 2, x_15); -lean_ctor_set(x_43, 3, x_12); +lean_ctor_set(x_43, 0, x_13); +lean_ctor_set(x_43, 1, x_15); +lean_ctor_set(x_43, 2, x_11); +lean_ctor_set(x_43, 3, x_10); lean_ctor_set(x_43, 4, x_42); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_13); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_11); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_12); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_14); 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); @@ -5501,25 +5501,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_54); -x_62 = l_Lean_FileMap_toPosition(x_54, x_53); -lean_dec(x_53); -x_63 = l_Lean_FileMap_toPosition(x_54, x_57); +lean_inc_ref(x_56); +x_62 = l_Lean_FileMap_toPosition(x_56, x_55); +lean_dec(x_55); +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_Lean_addTrace___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__0_spec__1___redArg___closed__1; -if (x_55 == 0) +if (x_51 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_62; -x_11 = x_51; -x_12 = x_65; +x_10 = x_65; +x_11 = x_64; +x_12 = x_53; x_13 = x_52; -x_14 = x_61; -x_15 = x_64; -x_16 = x_56; +x_14 = x_54; +x_15 = x_62; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -5536,7 +5536,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_56); +lean_dec_ref(x_52); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -5545,13 +5545,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_62; -x_11 = x_51; -x_12 = x_65; +x_10 = x_65; +x_11 = x_64; +x_12 = x_53; x_13 = x_52; -x_14 = x_61; -x_15 = x_64; -x_16 = x_56; +x_14 = x_54; +x_15 = x_62; +x_16 = x_61; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -5565,24 +5565,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_54); -x_69 = l_Lean_FileMap_toPosition(x_54, x_53); -lean_dec(x_53); -x_70 = l_Lean_FileMap_toPosition(x_54, x_57); +lean_inc_ref(x_56); +x_69 = l_Lean_FileMap_toPosition(x_56, x_55); +lean_dec(x_55); +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_Lean_addTrace___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__0_spec__1___redArg___closed__1; -if (x_55 == 0) +if (x_51 == 0) { lean_dec_ref(x_50); -x_10 = x_69; -x_11 = x_51; -x_12 = x_72; +x_10 = x_72; +x_11 = x_71; +x_12 = x_53; x_13 = x_52; -x_14 = x_68; -x_15 = x_71; -x_16 = x_56; +x_14 = x_54; +x_15 = x_69; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -5599,7 +5599,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_56); +lean_dec_ref(x_52); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -5608,13 +5608,13 @@ return x_75; } else { -x_10 = x_69; -x_11 = x_51; -x_12 = x_72; +x_10 = x_72; +x_11 = x_71; +x_12 = x_53; x_13 = x_52; -x_14 = x_68; -x_15 = x_71; -x_16 = x_56; +x_14 = x_54; +x_15 = x_69; +x_16 = x_68; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -5626,17 +5626,17 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_80, x_79); -lean_dec(x_80); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_81, x_78); +lean_dec(x_81); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_78; +x_51 = x_80; x_52 = x_79; -x_53 = x_84; -x_54 = x_81; -x_55 = x_82; +x_53 = x_78; +x_54 = x_82; +x_55 = x_84; x_56 = x_83; x_57 = x_84; goto block_76; @@ -5648,11 +5648,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_51 = x_80; x_52 = x_79; -x_53 = x_84; -x_54 = x_81; -x_55 = x_82; +x_53 = x_78; +x_54 = x_82; +x_55 = x_84; x_56 = x_83; x_57 = x_86; goto block_76; @@ -5661,19 +5661,19 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_92); -lean_dec(x_92); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_89); +x_95 = l_Lean_replaceRef(x_1, x_89); +lean_dec(x_89); +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_94; -x_79 = x_89; -x_80 = x_95; -x_81 = x_90; -x_82 = x_91; +x_78 = x_92; +x_79 = x_91; +x_80 = x_90; +x_81 = x_95; +x_82 = x_94; x_83 = x_93; x_84 = x_97; goto block_87; @@ -5685,11 +5685,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_94; -x_79 = x_89; -x_80 = x_95; -x_81 = x_90; -x_82 = x_91; +x_78 = x_92; +x_79 = x_91; +x_80 = x_90; +x_81 = x_95; +x_82 = x_94; x_83 = x_93; x_84 = x_98; goto block_87; @@ -5700,10 +5700,10 @@ goto block_87; if (x_107 == 0) { x_88 = x_104; -x_89 = x_106; -x_90 = x_101; +x_89 = x_101; +x_90 = x_103; x_91 = x_102; -x_92 = x_103; +x_92 = x_106; x_93 = x_105; x_94 = x_3; goto block_99; @@ -5711,10 +5711,10 @@ goto block_99; else { x_88 = x_104; -x_89 = x_106; -x_90 = x_101; +x_89 = x_101; +x_90 = x_103; x_91 = x_102; -x_92 = x_103; +x_92 = x_106; x_93 = x_105; x_94 = x_100; goto block_99; @@ -5739,14 +5739,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); -lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_114; -x_103 = x_113; +x_101 = x_113; +x_102 = x_110; +x_103 = x_114; x_104 = x_117; -x_105 = x_110; +x_105 = x_111; x_106 = x_109; x_107 = x_119; goto block_108; @@ -5756,14 +5756,14 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22_spec__22___redArg___closed__0; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__0_spec__11_spec__11_spec__11_spec__11(x_112, x_120); +lean_inc_ref(x_111); lean_inc_ref(x_110); lean_inc(x_113); -lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_114; -x_103 = x_113; +x_101 = x_113; +x_102 = x_110; +x_103 = x_114; x_104 = x_117; -x_105 = x_110; +x_105 = x_111; x_106 = x_109; x_107 = x_121; goto block_108; @@ -6039,7 +6039,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_ErrorExplanation_elabCheckedNamedError(lean_object* x_1, lean_object* x_2, lean_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; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; 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_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_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_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_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_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; +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_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; 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_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_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_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; x_172 = l___private_Lean_Elab_ErrorExplanation_0__Lean_Elab_ErrorExplanation_macroDeclMap; lean_inc(x_1); x_173 = l_Lean_Syntax_getKind(x_1); @@ -6275,25 +6275,25 @@ lean_dec(x_38); x_40 = lean_alloc_ctor(6, 2, 0); lean_ctor_set(x_40, 0, x_35); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19(x_40, x_31, x_28, x_26, x_27, x_33, x_29); +x_41 = l_Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19(x_40, x_34, x_31, x_33, x_30, x_26, x_32); x_42 = !lean_is_exclusive(x_41); if (x_42 == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_43 = lean_ctor_get(x_41, 0); lean_dec(x_43); -x_44 = l_Lean_Syntax_getId(x_32); +x_44 = l_Lean_Syntax_getId(x_27); x_45 = lean_erase_macro_scopes(x_44); lean_inc(x_45); -lean_inc(x_32); +lean_inc(x_27); x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_32); +lean_ctor_set(x_46, 0, x_27); lean_ctor_set(x_46, 1, x_45); lean_ctor_set_tag(x_41, 6); lean_ctor_set(x_41, 0, x_46); -x_47 = l_Lean_Elab_pushInfoLeaf___at___00Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19_spec__19(x_41, x_31, x_28, x_26, x_27, x_33, x_29); +x_47 = l_Lean_Elab_pushInfoLeaf___at___00Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19_spec__19(x_41, x_34, x_31, x_33, x_30, x_26, x_32); lean_dec_ref(x_47); -x_48 = lean_st_ref_get(x_29); +x_48 = lean_st_ref_get(x_32); x_49 = lean_ctor_get(x_48, 0); lean_inc_ref(x_49); lean_dec_ref(x_48); @@ -6310,17 +6310,17 @@ x_54 = l_Lean_Elab_ErrorExplanation_elabCheckedNamedError___closed__3; x_55 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_55, 0, x_53); lean_ctor_set(x_55, 1, x_54); -lean_inc_ref(x_33); -x_56 = l_Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22(x_32, x_55, x_31, x_28, x_26, x_27, x_33, x_29); -lean_dec(x_32); +lean_inc_ref(x_26); +x_56 = l_Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22(x_27, x_55, x_34, x_31, x_33, x_30, x_26, x_32); +lean_dec(x_27); lean_dec_ref(x_56); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6339,14 +6339,14 @@ lean_dec_ref(x_58); if (lean_obj_tag(x_59) == 0) { lean_dec(x_45); -lean_dec(x_32); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +lean_dec(x_27); +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6374,17 +6374,17 @@ x_68 = l_Lean_Elab_ErrorExplanation_elabCheckedNamedError___closed__9; x_69 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); -lean_inc_ref(x_33); -x_70 = l_Lean_logWarningAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__24(x_32, x_69, x_31, x_28, x_26, x_27, x_33, x_29); -lean_dec(x_32); +lean_inc_ref(x_26); +x_70 = l_Lean_logWarningAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__24(x_27, x_69, x_34, x_31, x_33, x_30, x_26, x_32); +lean_dec(x_27); lean_dec_ref(x_70); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6394,18 +6394,18 @@ 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_dec(x_41); -x_71 = l_Lean_Syntax_getId(x_32); +x_71 = l_Lean_Syntax_getId(x_27); x_72 = lean_erase_macro_scopes(x_71); lean_inc(x_72); -lean_inc(x_32); +lean_inc(x_27); x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_32); +lean_ctor_set(x_73, 0, x_27); lean_ctor_set(x_73, 1, x_72); x_74 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_74, 0, x_73); -x_75 = l_Lean_Elab_pushInfoLeaf___at___00Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19_spec__19(x_74, x_31, x_28, x_26, x_27, x_33, x_29); +x_75 = l_Lean_Elab_pushInfoLeaf___at___00Lean_Elab_addCompletionInfo___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__19_spec__19(x_74, x_34, x_31, x_33, x_30, x_26, x_32); lean_dec_ref(x_75); -x_76 = lean_st_ref_get(x_29); +x_76 = lean_st_ref_get(x_32); x_77 = lean_ctor_get(x_76, 0); lean_inc_ref(x_77); lean_dec_ref(x_76); @@ -6422,17 +6422,17 @@ x_82 = l_Lean_Elab_ErrorExplanation_elabCheckedNamedError___closed__3; x_83 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_83, 0, x_81); lean_ctor_set(x_83, 1, x_82); -lean_inc_ref(x_33); -x_84 = l_Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22(x_32, x_83, x_31, x_28, x_26, x_27, x_33, x_29); -lean_dec(x_32); +lean_inc_ref(x_26); +x_84 = l_Lean_logErrorAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__22(x_27, x_83, x_34, x_31, x_33, x_30, x_26, x_32); +lean_dec(x_27); lean_dec_ref(x_84); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6451,14 +6451,14 @@ lean_dec_ref(x_86); if (lean_obj_tag(x_87) == 0) { lean_dec(x_72); -lean_dec(x_32); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +lean_dec(x_27); +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6486,17 +6486,17 @@ x_96 = l_Lean_Elab_ErrorExplanation_elabCheckedNamedError___closed__9; x_97 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_97, 0, x_95); lean_ctor_set(x_97, 1, x_96); -lean_inc_ref(x_33); -x_98 = l_Lean_logWarningAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__24(x_32, x_97, x_31, x_28, x_26, x_27, x_33, x_29); -lean_dec(x_32); +lean_inc_ref(x_26); +x_98 = l_Lean_logWarningAt___at___00Lean_Elab_ErrorExplanation_elabCheckedNamedError_spec__24(x_27, x_97, x_34, x_31, x_33, x_30, x_26, x_32); +lean_dec(x_27); lean_dec_ref(x_98); -x_10 = x_30; -x_11 = x_31; -x_12 = x_28; -x_13 = x_26; -x_14 = x_27; -x_15 = x_33; -x_16 = x_29; +x_10 = x_29; +x_11 = x_34; +x_12 = x_31; +x_13 = x_33; +x_14 = x_30; +x_15 = x_26; +x_16 = x_32; x_17 = lean_box(0); goto block_25; } @@ -6513,13 +6513,13 @@ lean_inc(x_1); x_114 = l_Lean_Syntax_setArgs(x_1, x_113); x_26 = x_100; x_27 = x_101; -x_28 = x_102; -x_29 = x_104; +x_28 = lean_box(0); +x_29 = x_106; x_30 = x_105; -x_31 = x_106; -x_32 = x_108; -x_33 = x_107; -x_34 = lean_box(0); +x_31 = x_104; +x_32 = x_107; +x_33 = x_109; +x_34 = x_108; x_35 = x_114; goto block_99; } @@ -6534,14 +6534,14 @@ if (x_126 == 0) lean_dec(x_125); lean_inc(x_1); x_26 = x_116; -x_27 = x_117; -x_28 = x_118; -x_29 = x_119; -x_30 = x_127; -x_31 = x_120; -x_32 = x_123; -x_33 = x_121; -x_34 = lean_box(0); +x_27 = x_123; +x_28 = lean_box(0); +x_29 = x_127; +x_30 = x_119; +x_31 = x_118; +x_32 = x_120; +x_33 = x_122; +x_34 = x_121; x_35 = x_1; goto block_99; } @@ -6559,15 +6559,15 @@ if (x_133 == 0) { lean_dec(x_130); x_100 = x_116; -x_101 = x_117; -x_102 = x_118; +x_101 = x_123; +x_102 = lean_box(0); x_103 = x_128; -x_104 = x_119; -x_105 = x_127; -x_106 = x_120; -x_107 = x_121; -x_108 = x_123; -x_109 = lean_box(0); +x_104 = x_118; +x_105 = x_119; +x_106 = x_127; +x_107 = x_120; +x_108 = x_121; +x_109 = x_122; x_110 = x_131; x_111 = x_132; goto block_115; @@ -6576,15 +6576,15 @@ else { lean_dec(x_132); x_100 = x_116; -x_101 = x_117; -x_102 = x_118; +x_101 = x_123; +x_102 = lean_box(0); x_103 = x_128; -x_104 = x_119; -x_105 = x_127; -x_106 = x_120; -x_107 = x_121; -x_108 = x_123; -x_109 = lean_box(0); +x_104 = x_118; +x_105 = x_119; +x_106 = x_127; +x_107 = x_120; +x_108 = x_121; +x_109 = x_122; x_110 = x_131; x_111 = x_130; goto block_115; @@ -6598,12 +6598,12 @@ x_142 = lean_unsigned_to_nat(2u); x_143 = l_Lean_Syntax_getArg(x_1, x_142); x_144 = lean_unsigned_to_nat(5u); x_116 = x_135; -x_117 = x_136; -x_118 = x_137; -x_119 = x_138; +x_117 = lean_box(0); +x_118 = x_138; +x_119 = x_137; x_120 = x_139; -x_121 = x_140; -x_122 = lean_box(0); +x_121 = x_141; +x_122 = x_140; x_123 = x_143; x_124 = x_144; goto block_134; @@ -6623,12 +6623,12 @@ x_156 = lean_unsigned_to_nat(1u); x_157 = l_Lean_Syntax_getArg(x_1, x_156); x_158 = lean_unsigned_to_nat(4u); x_116 = x_146; -x_117 = x_147; -x_118 = x_148; -x_119 = x_149; +x_117 = lean_box(0); +x_118 = x_149; +x_119 = x_148; x_120 = x_150; -x_121 = x_151; -x_122 = lean_box(0); +x_121 = x_152; +x_122 = x_151; x_123 = x_157; x_124 = x_158; goto block_134; @@ -6636,24 +6636,24 @@ goto block_134; else { x_135 = x_146; -x_136 = x_147; +x_136 = lean_box(0); x_137 = x_148; x_138 = x_149; x_139 = x_150; x_140 = x_151; -x_141 = lean_box(0); +x_141 = x_152; goto block_145; } } else { x_135 = x_146; -x_136 = x_147; +x_136 = lean_box(0); x_137 = x_148; x_138 = x_149; x_139 = x_150; x_140 = x_151; -x_141 = lean_box(0); +x_141 = x_152; goto block_145; } } @@ -6669,25 +6669,25 @@ lean_object* x_169; uint8_t x_170; x_169 = l_Lean_Elab_ErrorExplanation_expandNamedErrorMacro___closed__9; lean_inc(x_1); x_170 = l_Lean_Syntax_isOfKind(x_1, x_169); -x_146 = x_162; -x_147 = x_163; -x_148 = x_161; -x_149 = x_165; -x_150 = x_160; -x_151 = x_164; -x_152 = lean_box(0); +x_146 = x_164; +x_147 = lean_box(0); +x_148 = x_163; +x_149 = x_161; +x_150 = x_165; +x_151 = x_162; +x_152 = x_160; x_153 = x_170; goto block_159; } else { -x_146 = x_162; -x_147 = x_163; -x_148 = x_161; -x_149 = x_165; -x_150 = x_160; -x_151 = x_164; -x_152 = lean_box(0); +x_146 = x_164; +x_147 = lean_box(0); +x_148 = x_163; +x_149 = x_161; +x_150 = x_165; +x_151 = x_162; +x_152 = x_160; x_153 = x_168; goto block_159; } @@ -9472,7 +9472,7 @@ return x_91; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_105; uint8_t x_106; +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_105; uint8_t x_106; x_92 = lean_unsigned_to_nat(2u); x_93 = l_Lean_Syntax_getArg(x_1, x_92); x_105 = l_Lean_Elab_ErrorExplanation_expandNamedErrorMacro___closed__54; @@ -10186,12 +10186,12 @@ lean_dec_ref(x_118); x_254 = l_Lean_Syntax_getPos_x3f(x_93, x_132); if (lean_obj_tag(x_254) == 0) { -x_94 = x_253; -x_95 = x_123; -x_96 = x_116; +x_94 = x_116; +x_95 = x_119; +x_96 = x_123; x_97 = lean_box(0); -x_98 = x_119; -x_99 = x_132; +x_98 = x_132; +x_99 = x_253; x_100 = x_117; x_101 = x_87; goto block_104; @@ -10202,12 +10202,12 @@ lean_object* x_255; x_255 = lean_ctor_get(x_254, 0); lean_inc(x_255); lean_dec_ref(x_254); -x_94 = x_253; -x_95 = x_123; -x_96 = x_116; +x_94 = x_116; +x_95 = x_119; +x_96 = x_123; x_97 = lean_box(0); -x_98 = x_119; -x_99 = x_132; +x_98 = x_132; +x_99 = x_253; x_100 = x_117; x_101 = x_255; goto block_104; @@ -10667,12 +10667,12 @@ lean_dec_ref(x_322); x_368 = l_Lean_Syntax_getPos_x3f(x_93, x_336); if (lean_obj_tag(x_368) == 0) { -x_94 = x_367; -x_95 = x_327; -x_96 = x_320; +x_94 = x_320; +x_95 = x_323; +x_96 = x_327; x_97 = lean_box(0); -x_98 = x_323; -x_99 = x_336; +x_98 = x_336; +x_99 = x_367; x_100 = x_321; x_101 = x_87; goto block_104; @@ -10683,12 +10683,12 @@ lean_object* x_369; x_369 = lean_ctor_get(x_368, 0); lean_inc(x_369); lean_dec_ref(x_368); -x_94 = x_367; -x_95 = x_327; -x_96 = x_320; +x_94 = x_320; +x_95 = x_323; +x_96 = x_327; x_97 = lean_box(0); -x_98 = x_323; -x_99 = x_336; +x_98 = x_336; +x_99 = x_367; x_100 = x_321; x_101 = x_369; goto block_104; @@ -10780,18 +10780,18 @@ return x_424; block_104: { lean_object* x_102; -x_102 = l_Lean_Syntax_getTailPos_x3f(x_93, x_99); +x_102 = l_Lean_Syntax_getTailPos_x3f(x_93, x_98); lean_dec(x_93); if (lean_obj_tag(x_102) == 0) { lean_inc(x_101); -x_5 = x_95; -x_6 = x_94; -x_7 = x_96; +x_5 = x_94; +x_6 = x_96; +x_7 = x_95; x_8 = lean_box(0); x_9 = x_101; -x_10 = x_98; -x_11 = x_100; +x_10 = x_100; +x_11 = x_99; x_12 = x_101; goto block_83; } @@ -10801,13 +10801,13 @@ lean_object* x_103; x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); lean_dec_ref(x_102); -x_5 = x_95; -x_6 = x_94; -x_7 = x_96; +x_5 = x_94; +x_6 = x_96; +x_7 = x_95; x_8 = lean_box(0); x_9 = x_101; -x_10 = x_98; -x_11 = x_100; +x_10 = x_100; +x_11 = x_99; x_12 = x_103; goto block_83; } @@ -10817,13 +10817,13 @@ goto block_83; block_83: { lean_object* x_13; uint8_t x_14; -x_13 = l_Lean_getMainModule___at___00Lean_Elab_ErrorExplanation_elabRegisterErrorExplanation_spec__14___redArg(x_10); +x_13 = l_Lean_getMainModule___at___00Lean_Elab_ErrorExplanation_elabRegisterErrorExplanation_spec__14___redArg(x_7); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_st_ref_take(x_10); +x_16 = lean_st_ref_take(x_7); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -10835,7 +10835,7 @@ lean_inc_ref(x_20); x_21 = lean_ctor_get(x_20, 2); lean_inc(x_21); lean_dec_ref(x_20); -x_22 = l_Lean_DeclarationRange_ofStringPositions(x_6, x_9, x_12); +x_22 = l_Lean_DeclarationRange_ofStringPositions(x_11, x_9, x_12); lean_dec(x_12); lean_dec(x_9); x_23 = lean_alloc_ctor(0, 2, 0); @@ -10844,17 +10844,17 @@ lean_ctor_set(x_23, 1, x_22); x_24 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_24, 0, x_23); x_25 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_25, 0, x_5); -lean_ctor_set(x_25, 1, x_11); +lean_ctor_set(x_25, 0, x_6); +lean_ctor_set(x_25, 1, x_10); lean_ctor_set(x_25, 2, x_24); x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 0, x_5); lean_ctor_set(x_26, 1, x_25); x_27 = lean_box(0); x_28 = l_Lean_PersistentEnvExtension_addEntry___redArg(x_19, x_18, x_26, x_21, x_27); lean_dec(x_21); lean_ctor_set(x_16, 0, x_28); -x_29 = lean_st_ref_set(x_10, x_16); +x_29 = lean_st_ref_set(x_7, x_16); x_30 = lean_box(0); lean_ctor_set(x_13, 0, x_30); return x_13; @@ -10891,7 +10891,7 @@ lean_inc_ref(x_43); x_44 = lean_ctor_get(x_43, 2); lean_inc(x_44); lean_dec_ref(x_43); -x_45 = l_Lean_DeclarationRange_ofStringPositions(x_6, x_9, x_12); +x_45 = l_Lean_DeclarationRange_ofStringPositions(x_11, x_9, x_12); lean_dec(x_12); lean_dec(x_9); x_46 = lean_alloc_ctor(0, 2, 0); @@ -10900,11 +10900,11 @@ lean_ctor_set(x_46, 1, x_45); x_47 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_47, 0, x_46); x_48 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_48, 0, x_5); -lean_ctor_set(x_48, 1, x_11); +lean_ctor_set(x_48, 0, x_6); +lean_ctor_set(x_48, 1, x_10); lean_ctor_set(x_48, 2, x_47); x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_7); +lean_ctor_set(x_49, 0, x_5); lean_ctor_set(x_49, 1, x_48); x_50 = lean_box(0); x_51 = l_Lean_PersistentEnvExtension_addEntry___redArg(x_42, x_31, x_49, x_44, x_50); @@ -10921,7 +10921,7 @@ lean_ctor_set(x_52, 7, x_38); lean_ctor_set(x_52, 8, x_39); lean_ctor_set(x_52, 9, x_40); lean_ctor_set(x_52, 10, x_41); -x_53 = lean_st_ref_set(x_10, x_52); +x_53 = lean_st_ref_set(x_7, x_52); x_54 = lean_box(0); lean_ctor_set(x_13, 0, x_54); return x_13; @@ -10933,7 +10933,7 @@ lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean x_55 = lean_ctor_get(x_13, 0); lean_inc(x_55); lean_dec(x_13); -x_56 = lean_st_ref_take(x_10); +x_56 = lean_st_ref_take(x_7); x_57 = lean_ctor_get(x_56, 0); lean_inc_ref(x_57); x_58 = lean_ctor_get(x_56, 1); @@ -10979,7 +10979,7 @@ lean_inc_ref(x_70); x_71 = lean_ctor_get(x_70, 2); lean_inc(x_71); lean_dec_ref(x_70); -x_72 = l_Lean_DeclarationRange_ofStringPositions(x_6, x_9, x_12); +x_72 = l_Lean_DeclarationRange_ofStringPositions(x_11, x_9, x_12); lean_dec(x_12); lean_dec(x_9); x_73 = lean_alloc_ctor(0, 2, 0); @@ -10988,11 +10988,11 @@ lean_ctor_set(x_73, 1, x_72); x_74 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_74, 0, x_73); x_75 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_75, 0, x_5); -lean_ctor_set(x_75, 1, x_11); +lean_ctor_set(x_75, 0, x_6); +lean_ctor_set(x_75, 1, x_10); lean_ctor_set(x_75, 2, x_74); x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_7); +lean_ctor_set(x_76, 0, x_5); lean_ctor_set(x_76, 1, x_75); x_77 = lean_box(0); x_78 = l_Lean_PersistentEnvExtension_addEntry___redArg(x_69, x_57, x_76, x_71, x_77); @@ -11013,7 +11013,7 @@ lean_ctor_set(x_79, 7, x_64); lean_ctor_set(x_79, 8, x_65); lean_ctor_set(x_79, 9, x_66); lean_ctor_set(x_79, 10, x_67); -x_80 = lean_st_ref_set(x_10, x_79); +x_80 = lean_st_ref_set(x_7, x_79); x_81 = lean_box(0); x_82 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_82, 0, x_81); diff --git a/stage0/stdlib/Lean/Elab/Macro.c b/stage0/stdlib/Lean/Elab/Macro.c index 4b1ef0848867..ab0aef6d3992 100644 --- a/stage0/stdlib/Lean/Elab/Macro.c +++ b/stage0/stdlib/Lean/Elab/Macro.c @@ -3855,7 +3855,7 @@ return x_179; } else { -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_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; uint8_t x_286; uint8_t x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; size_t x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; 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_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; size_t 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; uint8_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; size_t x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; size_t x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_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; size_t x_487; lean_object* x_488; lean_object* x_489; uint8_t x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; uint8_t x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; uint8_t x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_730; uint8_t x_731; +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_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; uint8_t x_286; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; uint8_t x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; size_t x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; size_t 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_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; uint8_t x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; size_t x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; uint8_t x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; size_t 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_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; size_t x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; uint8_t x_499; lean_object* x_500; lean_object* x_501; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; uint8_t x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; uint8_t x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; uint8_t x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_730; uint8_t x_731; x_180 = lean_unsigned_to_nat(0u); x_730 = l_Lean_Syntax_getArg(x_1, x_180); x_731 = l_Lean_Syntax_isNone(x_730); @@ -3920,16 +3920,16 @@ goto block_729; block_225: { lean_object* x_195; lean_object* x_196; uint8_t x_197; -x_195 = l_Lean_Syntax_getArgs(x_182); +x_195 = l_Lean_Syntax_getArgs(x_184); x_196 = lean_array_get_size(x_195); lean_dec_ref(x_195); -x_197 = lean_nat_dec_eq(x_196, x_181); +x_197 = lean_nat_dec_eq(x_196, x_182); lean_dec(x_196); if (x_197 == 0) { lean_object* x_198; +lean_dec_ref(x_187); lean_dec(x_185); -lean_dec_ref(x_184); x_198 = l_Lean_Elab_Command_elabMacro___lam__0(x_192, x_193); if (lean_obj_tag(x_198) == 0) { @@ -3943,40 +3943,40 @@ if (lean_obj_tag(x_200) == 0) lean_object* x_201; lean_object* x_202; lean_dec_ref(x_200); x_201 = lean_ctor_get(x_192, 5); -x_202 = l_Lean_Syntax_getArg(x_182, x_181); -lean_dec(x_182); +x_202 = l_Lean_Syntax_getArg(x_184, x_182); +lean_dec(x_184); if (lean_obj_tag(x_201) == 0) { lean_object* x_203; x_203 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabMacro_spec__1_spec__18___redArg(x_193); lean_dec_ref(x_203); -x_48 = x_193; +x_48 = x_181; x_49 = x_199; -x_50 = x_202; -x_51 = x_183; -x_52 = x_191; -x_53 = x_187; -x_54 = x_186; -x_55 = x_189; -x_56 = x_192; -x_57 = x_188; -x_58 = x_190; +x_50 = x_183; +x_51 = x_202; +x_52 = x_186; +x_53 = x_193; +x_54 = x_192; +x_55 = x_188; +x_56 = x_190; +x_57 = x_191; +x_58 = x_189; x_59 = lean_box(0); goto block_65; } else { -x_48 = x_193; +x_48 = x_181; x_49 = x_199; -x_50 = x_202; -x_51 = x_183; -x_52 = x_191; -x_53 = x_187; -x_54 = x_186; -x_55 = x_189; -x_56 = x_192; -x_57 = x_188; -x_58 = x_190; +x_50 = x_183; +x_51 = x_202; +x_52 = x_186; +x_53 = x_193; +x_54 = x_192; +x_55 = x_188; +x_56 = x_190; +x_57 = x_191; +x_58 = x_189; x_59 = lean_box(0); goto block_65; } @@ -3989,12 +3989,12 @@ lean_dec(x_193); lean_dec_ref(x_192); lean_dec(x_191); lean_dec(x_190); -lean_dec_ref(x_189); +lean_dec(x_189); lean_dec_ref(x_188); -lean_dec(x_187); lean_dec_ref(x_186); +lean_dec(x_184); lean_dec(x_183); -lean_dec(x_182); +lean_dec_ref(x_181); x_204 = !lean_is_exclusive(x_200); if (x_204 == 0) { @@ -4019,12 +4019,12 @@ lean_dec(x_193); lean_dec_ref(x_192); lean_dec(x_191); lean_dec(x_190); -lean_dec_ref(x_189); +lean_dec(x_189); lean_dec_ref(x_188); -lean_dec(x_187); lean_dec_ref(x_186); +lean_dec(x_184); lean_dec(x_183); -lean_dec(x_182); +lean_dec_ref(x_181); x_207 = !lean_is_exclusive(x_198); if (x_207 == 0) { @@ -4060,8 +4060,8 @@ x_213 = lean_ctor_get(x_212, 0); lean_inc(x_213); lean_dec_ref(x_212); x_214 = lean_ctor_get(x_192, 5); -x_215 = l_Lean_Syntax_getArg(x_182, x_180); -lean_dec(x_182); +x_215 = l_Lean_Syntax_getArg(x_184, x_180); +lean_dec(x_184); if (lean_obj_tag(x_214) == 0) { lean_object* x_216; lean_object* x_217; @@ -4069,20 +4069,20 @@ x_216 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Co x_217 = lean_ctor_get(x_216, 0); lean_inc(x_217); lean_dec_ref(x_216); -x_155 = x_184; -x_156 = x_187; -x_157 = x_192; -x_158 = x_188; -x_159 = x_213; -x_160 = x_193; -x_161 = x_183; -x_162 = x_191; -x_163 = x_185; -x_164 = x_186; -x_165 = x_211; -x_166 = x_189; -x_167 = x_215; -x_168 = x_190; +x_155 = x_181; +x_156 = x_186; +x_157 = x_193; +x_158 = x_192; +x_159 = x_188; +x_160 = x_190; +x_161 = x_189; +x_162 = x_211; +x_163 = x_213; +x_164 = x_215; +x_165 = x_183; +x_166 = x_185; +x_167 = x_187; +x_168 = x_191; x_169 = x_217; x_170 = lean_box(0); goto block_176; @@ -4092,20 +4092,20 @@ else lean_object* x_218; x_218 = lean_ctor_get(x_214, 0); lean_inc(x_218); -x_155 = x_184; -x_156 = x_187; -x_157 = x_192; -x_158 = x_188; -x_159 = x_213; -x_160 = x_193; -x_161 = x_183; -x_162 = x_191; -x_163 = x_185; -x_164 = x_186; -x_165 = x_211; -x_166 = x_189; -x_167 = x_215; -x_168 = x_190; +x_155 = x_181; +x_156 = x_186; +x_157 = x_193; +x_158 = x_192; +x_159 = x_188; +x_160 = x_190; +x_161 = x_189; +x_162 = x_211; +x_163 = x_213; +x_164 = x_215; +x_165 = x_183; +x_166 = x_185; +x_167 = x_187; +x_168 = x_191; x_169 = x_218; x_170 = lean_box(0); goto block_176; @@ -4119,14 +4119,14 @@ lean_dec(x_193); lean_dec_ref(x_192); lean_dec(x_191); lean_dec(x_190); -lean_dec_ref(x_189); +lean_dec(x_189); lean_dec_ref(x_188); -lean_dec(x_187); +lean_dec_ref(x_187); lean_dec_ref(x_186); lean_dec(x_185); -lean_dec_ref(x_184); +lean_dec(x_184); lean_dec(x_183); -lean_dec(x_182); +lean_dec_ref(x_181); x_219 = !lean_is_exclusive(x_212); if (x_219 == 0) { @@ -4151,14 +4151,14 @@ lean_dec(x_193); lean_dec_ref(x_192); lean_dec(x_191); lean_dec(x_190); -lean_dec_ref(x_189); +lean_dec(x_189); lean_dec_ref(x_188); -lean_dec(x_187); +lean_dec_ref(x_187); lean_dec_ref(x_186); lean_dec(x_185); -lean_dec_ref(x_184); +lean_dec(x_184); lean_dec(x_183); -lean_dec(x_182); +lean_dec_ref(x_181); x_222 = !lean_is_exclusive(x_210); if (x_222 == 0) { @@ -4180,60 +4180,60 @@ return x_224; block_243: { lean_object* x_241; lean_object* x_242; -lean_inc_ref(x_234); -lean_inc(x_232); -lean_inc(x_236); +lean_inc_ref(x_229); +lean_inc(x_235); +lean_inc(x_227); x_241 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_241, 0, x_236); -lean_ctor_set(x_241, 1, x_232); -lean_ctor_set(x_241, 2, x_234); -x_242 = l_Lean_Syntax_node1(x_236, x_235, x_241); +lean_ctor_set(x_241, 0, x_227); +lean_ctor_set(x_241, 1, x_235); +lean_ctor_set(x_241, 2, x_229); +x_242 = l_Lean_Syntax_node1(x_227, x_232, x_241); x_181 = x_226; -x_182 = x_227; -x_183 = x_232; -x_184 = x_228; -x_185 = x_233; -x_186 = x_234; -x_187 = x_229; -x_188 = x_230; -x_189 = x_238; -x_190 = x_239; +x_182 = x_228; +x_183 = x_236; +x_184 = x_237; +x_185 = x_238; +x_186 = x_229; +x_187 = x_239; +x_188 = x_233; +x_189 = x_235; +x_190 = x_234; x_191 = x_242; -x_192 = x_231; -x_193 = x_237; +x_192 = x_230; +x_193 = x_231; x_194 = lean_box(0); goto block_225; } block_269: { lean_object* x_259; -x_259 = l_Lean_Elab_Command_elabMacro___lam__0(x_250, x_256); +x_259 = l_Lean_Elab_Command_elabMacro___lam__0(x_248, x_249); if (lean_obj_tag(x_259) == 0) { lean_object* x_260; lean_object* x_261; x_260 = lean_ctor_get(x_259, 0); lean_inc(x_260); lean_dec_ref(x_259); -x_261 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_250); +x_261 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_248); if (lean_obj_tag(x_261) == 0) { lean_dec_ref(x_261); -if (lean_obj_tag(x_253) == 0) +if (lean_obj_tag(x_254) == 0) { lean_object* x_262; -x_262 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabMacro_spec__1_spec__18___redArg(x_256); +x_262 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabMacro_spec__1_spec__18___redArg(x_249); lean_dec_ref(x_262); x_226 = x_244; -x_227 = x_245; -x_228 = x_247; -x_229 = x_248; -x_230 = x_249; -x_231 = x_250; -x_232 = x_251; -x_233 = x_252; -x_234 = x_254; -x_235 = x_255; -x_236 = x_260; +x_227 = x_260; +x_228 = x_246; +x_229 = x_247; +x_230 = x_248; +x_231 = x_249; +x_232 = x_250; +x_233 = x_251; +x_234 = x_252; +x_235 = x_253; +x_236 = x_255; x_237 = x_256; x_238 = x_257; x_239 = x_258; @@ -4242,18 +4242,18 @@ goto block_243; } else { -lean_dec_ref(x_253); +lean_dec_ref(x_254); x_226 = x_244; -x_227 = x_245; -x_228 = x_247; -x_229 = x_248; -x_230 = x_249; -x_231 = x_250; -x_232 = x_251; -x_233 = x_252; -x_234 = x_254; -x_235 = x_255; -x_236 = x_260; +x_227 = x_260; +x_228 = x_246; +x_229 = x_247; +x_230 = x_248; +x_231 = x_249; +x_232 = x_250; +x_233 = x_251; +x_234 = x_252; +x_235 = x_253; +x_236 = x_255; x_237 = x_256; x_238 = x_257; x_239 = x_258; @@ -4265,19 +4265,19 @@ else { uint8_t x_263; lean_dec(x_260); -lean_dec(x_258); -lean_dec_ref(x_257); +lean_dec_ref(x_258); +lean_dec(x_257); lean_dec(x_256); lean_dec(x_255); -lean_dec_ref(x_254); +lean_dec(x_254); lean_dec(x_253); lean_dec(x_252); -lean_dec(x_251); -lean_dec_ref(x_250); -lean_dec_ref(x_249); -lean_dec(x_248); +lean_dec_ref(x_251); +lean_dec(x_250); +lean_dec(x_249); +lean_dec_ref(x_248); lean_dec_ref(x_247); -lean_dec(x_245); +lean_dec_ref(x_244); x_263 = !lean_is_exclusive(x_261); if (x_263 == 0) { @@ -4298,19 +4298,19 @@ return x_265; else { uint8_t x_266; -lean_dec(x_258); -lean_dec_ref(x_257); +lean_dec_ref(x_258); +lean_dec(x_257); lean_dec(x_256); lean_dec(x_255); -lean_dec_ref(x_254); +lean_dec(x_254); lean_dec(x_253); lean_dec(x_252); -lean_dec(x_251); -lean_dec_ref(x_250); -lean_dec_ref(x_249); -lean_dec(x_248); +lean_dec_ref(x_251); +lean_dec(x_250); +lean_dec(x_249); +lean_dec_ref(x_248); lean_dec_ref(x_247); -lean_dec(x_245); +lean_dec_ref(x_244); x_266 = !lean_is_exclusive(x_259); if (x_266 == 0) { @@ -4333,38 +4333,38 @@ return x_268; if (x_286 == 0) { lean_dec(x_281); -lean_dec(x_280); +lean_dec(x_277); x_181 = x_270; -x_182 = x_271; -x_183 = x_278; -x_184 = x_273; -x_185 = x_279; -x_186 = x_282; -x_187 = x_274; -x_188 = x_275; -x_189 = x_284; -x_190 = x_285; -x_191 = x_277; -x_192 = x_276; -x_193 = x_283; +x_182 = x_272; +x_183 = x_282; +x_184 = x_283; +x_185 = x_284; +x_186 = x_274; +x_187 = x_285; +x_188 = x_278; +x_189 = x_280; +x_190 = x_279; +x_191 = x_273; +x_192 = x_275; +x_193 = x_276; x_194 = lean_box(0); goto block_225; } else { -lean_dec(x_277); +lean_dec(x_273); x_244 = x_270; -x_245 = x_271; -x_246 = lean_box(0); -x_247 = x_273; -x_248 = x_274; -x_249 = x_275; -x_250 = x_276; +x_245 = lean_box(0); +x_246 = x_272; +x_247 = x_274; +x_248 = x_275; +x_249 = x_276; +x_250 = x_277; x_251 = x_278; x_252 = x_279; -x_253 = x_281; -x_254 = x_282; -x_255 = x_280; +x_253 = x_280; +x_254 = x_281; +x_255 = x_282; x_256 = x_283; x_257 = x_284; x_258 = x_285; @@ -4374,80 +4374,80 @@ goto block_269; block_359: { lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; size_t x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; -lean_inc_ref(x_310); -x_314 = l_Array_append___redArg(x_310, x_313); +lean_inc_ref(x_290); +x_314 = l_Array_append___redArg(x_290, x_313); lean_dec_ref(x_313); +lean_inc(x_306); lean_inc(x_307); -lean_inc(x_297); x_315 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_315, 0, x_297); -lean_ctor_set(x_315, 1, x_307); +lean_ctor_set(x_315, 0, x_307); +lean_ctor_set(x_315, 1, x_306); lean_ctor_set(x_315, 2, x_314); x_316 = l_Lean_Elab_Command_elabMacro___closed__47; x_317 = l_Lean_Elab_Command_elabMacro___closed__48; -lean_inc(x_297); +lean_inc(x_307); x_318 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_318, 0, x_297); +lean_ctor_set(x_318, 0, x_307); lean_ctor_set(x_318, 1, x_317); x_319 = l_Lean_Elab_Command_elabMacro___closed__49; -lean_inc(x_297); +lean_inc(x_307); x_320 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_320, 0, x_297); +lean_ctor_set(x_320, 0, x_307); lean_ctor_set(x_320, 1, x_319); x_321 = l_Lean_Elab_Command_elabMacro___closed__50; -lean_inc(x_297); +lean_inc(x_307); x_322 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_322, 0, x_297); +lean_ctor_set(x_322, 0, x_307); lean_ctor_set(x_322, 1, x_321); -x_323 = l_Nat_reprFast(x_309); -lean_inc(x_306); -x_324 = l_Lean_Syntax_mkNumLit(x_323, x_306); +x_323 = l_Nat_reprFast(x_294); +lean_inc(x_288); +x_324 = l_Lean_Syntax_mkNumLit(x_323, x_288); x_325 = l_Lean_Elab_Command_elabMacro___closed__51; -lean_inc(x_297); +lean_inc(x_307); x_326 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_326, 0, x_297); +lean_ctor_set(x_326, 0, x_307); lean_ctor_set(x_326, 1, x_325); -lean_inc(x_297); -x_327 = l_Lean_Syntax_node5(x_297, x_316, x_318, x_320, x_322, x_324, x_326); lean_inc(x_307); -lean_inc(x_297); -x_328 = l_Lean_Syntax_node1(x_297, x_307, x_327); -x_329 = lean_array_size(x_299); -x_330 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacro_spec__21(x_329, x_303, x_299); -lean_inc_ref(x_310); -x_331 = l_Array_append___redArg(x_310, x_330); +x_327 = l_Lean_Syntax_node5(x_307, x_316, x_318, x_320, x_322, x_324, x_326); +lean_inc(x_306); +lean_inc(x_307); +x_328 = l_Lean_Syntax_node1(x_307, x_306, x_327); +x_329 = lean_array_size(x_309); +x_330 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacro_spec__21(x_329, x_303, x_309); +lean_inc_ref(x_290); +x_331 = l_Array_append___redArg(x_290, x_330); lean_dec_ref(x_330); +lean_inc(x_306); lean_inc(x_307); -lean_inc(x_297); x_332 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_332, 0, x_297); -lean_ctor_set(x_332, 1, x_307); +lean_ctor_set(x_332, 0, x_307); +lean_ctor_set(x_332, 1, x_306); lean_ctor_set(x_332, 2, x_331); x_333 = l_Lean_Elab_Command_elabMacro___closed__52; -lean_inc(x_297); +lean_inc(x_307); x_334 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_334, 0, x_297); +lean_ctor_set(x_334, 0, x_307); lean_ctor_set(x_334, 1, x_333); x_335 = l_Lean_Elab_Command_elabMacro___closed__53; -x_336 = lean_array_push(x_335, x_296); -x_337 = lean_array_push(x_336, x_291); -lean_inc(x_290); -x_338 = lean_array_push(x_337, x_290); -x_339 = lean_array_push(x_338, x_295); -x_340 = lean_array_push(x_339, x_293); +x_336 = lean_array_push(x_335, x_298); +x_337 = lean_array_push(x_336, x_312); +lean_inc(x_300); +x_338 = lean_array_push(x_337, x_300); +x_339 = lean_array_push(x_338, x_308); +x_340 = lean_array_push(x_339, x_296); x_341 = lean_array_push(x_340, x_315); x_342 = lean_array_push(x_341, x_328); x_343 = lean_array_push(x_342, x_332); x_344 = lean_array_push(x_343, x_334); -lean_inc(x_292); -x_345 = lean_array_push(x_344, x_292); +lean_inc(x_311); +x_345 = lean_array_push(x_344, x_311); x_346 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_346, 0, x_297); -lean_ctor_set(x_346, 1, x_308); +lean_ctor_set(x_346, 0, x_307); +lean_ctor_set(x_346, 1, x_304); lean_ctor_set(x_346, 2, x_345); -lean_inc(x_312); -lean_inc_ref(x_289); -x_347 = l_Lean_Elab_Command_elabSyntax(x_346, x_289, x_312); +lean_inc(x_291); +lean_inc_ref(x_301); +x_347 = l_Lean_Elab_Command_elabSyntax(x_346, x_301, x_291); if (lean_obj_tag(x_347) == 0) { lean_object* x_348; lean_object* x_349; @@ -4455,54 +4455,54 @@ x_348 = lean_ctor_get(x_347, 0); lean_inc(x_348); lean_dec_ref(x_347); x_349 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_349, 0, x_306); +lean_ctor_set(x_349, 0, x_288); lean_ctor_set(x_349, 1, x_348); -lean_ctor_set(x_349, 2, x_302); -if (x_288 == 0) -{ -lean_dec(x_311); -lean_dec(x_294); -x_181 = x_300; -x_182 = x_301; -x_183 = x_307; -x_184 = x_317; -x_185 = x_292; -x_186 = x_310; -x_187 = x_304; -x_188 = x_305; -x_189 = x_325; -x_190 = x_349; -x_191 = x_290; -x_192 = x_289; -x_193 = x_312; +lean_ctor_set(x_349, 2, x_295); +if (x_299 == 0) +{ +lean_dec(x_305); +lean_dec(x_297); +x_181 = x_325; +x_182 = x_289; +x_183 = x_349; +x_184 = x_310; +x_185 = x_311; +x_186 = x_290; +x_187 = x_317; +x_188 = x_292; +x_189 = x_306; +x_190 = x_293; +x_191 = x_300; +x_192 = x_301; +x_193 = x_291; x_194 = lean_box(0); goto block_225; } else { lean_object* x_350; uint8_t x_351; -x_350 = l_Lean_Syntax_getArg(x_290, x_180); +x_350 = l_Lean_Syntax_getArg(x_300, x_180); lean_inc(x_350); -x_351 = l_Lean_Syntax_matchesNull(x_350, x_300); +x_351 = l_Lean_Syntax_matchesNull(x_350, x_289); if (x_351 == 0) { lean_dec(x_350); -x_270 = x_300; -x_271 = x_301; -x_272 = lean_box(0); -x_273 = x_317; -x_274 = x_304; -x_275 = x_305; -x_276 = x_289; -x_277 = x_290; -x_278 = x_307; -x_279 = x_292; -x_280 = x_294; -x_281 = x_311; -x_282 = x_310; -x_283 = x_312; -x_284 = x_325; -x_285 = x_349; +x_270 = x_325; +x_271 = lean_box(0); +x_272 = x_289; +x_273 = x_300; +x_274 = x_290; +x_275 = x_301; +x_276 = x_291; +x_277 = x_305; +x_278 = x_292; +x_279 = x_293; +x_280 = x_306; +x_281 = x_297; +x_282 = x_349; +x_283 = x_310; +x_284 = x_311; +x_285 = x_317; x_286 = x_351; goto block_287; } @@ -4512,49 +4512,49 @@ lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; x_352 = l_Lean_Syntax_getArg(x_350, x_180); lean_dec(x_350); x_353 = l_Lean_Elab_Command_elabMacro___closed__54; -lean_inc_ref(x_305); -x_354 = l_Lean_Name_mkStr4(x_5, x_6, x_305, x_353); +lean_inc_ref(x_292); +x_354 = l_Lean_Name_mkStr4(x_5, x_6, x_292, x_353); x_355 = l_Lean_Syntax_isOfKind(x_352, x_354); lean_dec(x_354); if (x_355 == 0) { -x_270 = x_300; -x_271 = x_301; -x_272 = lean_box(0); -x_273 = x_317; -x_274 = x_304; -x_275 = x_305; -x_276 = x_289; -x_277 = x_290; -x_278 = x_307; -x_279 = x_292; -x_280 = x_294; -x_281 = x_311; -x_282 = x_310; -x_283 = x_312; -x_284 = x_325; -x_285 = x_349; +x_270 = x_325; +x_271 = lean_box(0); +x_272 = x_289; +x_273 = x_300; +x_274 = x_290; +x_275 = x_301; +x_276 = x_291; +x_277 = x_305; +x_278 = x_292; +x_279 = x_293; +x_280 = x_306; +x_281 = x_297; +x_282 = x_349; +x_283 = x_310; +x_284 = x_311; +x_285 = x_317; x_286 = x_355; goto block_287; } else { -lean_dec(x_290); -x_244 = x_300; -x_245 = x_301; -x_246 = lean_box(0); -x_247 = x_317; -x_248 = x_304; -x_249 = x_305; -x_250 = x_289; -x_251 = x_307; -x_252 = x_292; -x_253 = x_311; -x_254 = x_310; -x_255 = x_294; -x_256 = x_312; -x_257 = x_325; -x_258 = x_349; +lean_dec(x_300); +x_244 = x_325; +x_245 = lean_box(0); +x_246 = x_289; +x_247 = x_290; +x_248 = x_301; +x_249 = x_291; +x_250 = x_305; +x_251 = x_292; +x_252 = x_293; +x_253 = x_306; +x_254 = x_297; +x_255 = x_349; +x_256 = x_310; +x_257 = x_311; +x_258 = x_317; goto block_269; } } @@ -4563,19 +4563,19 @@ goto block_269; else { uint8_t x_356; -lean_dec(x_312); lean_dec(x_311); -lean_dec_ref(x_310); -lean_dec(x_307); +lean_dec(x_310); lean_dec(x_306); -lean_dec_ref(x_305); -lean_dec(x_304); -lean_dec_ref(x_302); -lean_dec(x_301); -lean_dec(x_294); -lean_dec(x_292); -lean_dec(x_290); -lean_dec_ref(x_289); +lean_dec(x_305); +lean_dec_ref(x_301); +lean_dec(x_300); +lean_dec(x_297); +lean_dec_ref(x_295); +lean_dec(x_293); +lean_dec_ref(x_292); +lean_dec(x_291); +lean_dec_ref(x_290); +lean_dec(x_288); x_356 = !lean_is_exclusive(x_347); if (x_356 == 0) { @@ -4596,39 +4596,39 @@ return x_358; block_401: { lean_object* x_386; lean_object* x_387; -lean_inc_ref(x_382); -x_386 = l_Array_append___redArg(x_382, x_385); +lean_inc_ref(x_363); +x_386 = l_Array_append___redArg(x_363, x_385); lean_dec_ref(x_385); lean_inc(x_378); -lean_inc(x_369); +lean_inc(x_379); x_387 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_387, 0, x_369); +lean_ctor_set(x_387, 0, x_379); lean_ctor_set(x_387, 1, x_378); lean_ctor_set(x_387, 2, x_386); -if (lean_obj_tag(x_362) == 0) +if (lean_obj_tag(x_360) == 0) { lean_object* x_388; x_388 = l_Lean_Elab_Command_elabMacro___closed__11; -x_288 = x_360; -x_289 = x_361; -x_290 = x_364; -x_291 = x_363; +x_288 = x_361; +x_289 = x_362; +x_290 = x_363; +x_291 = x_364; x_292 = x_365; -x_293 = x_387; -x_294 = x_366; +x_293 = x_366; +x_294 = x_367; x_295 = x_368; -x_296 = x_367; +x_296 = x_387; x_297 = x_369; -x_298 = lean_box(0); +x_298 = x_370; x_299 = x_371; x_300 = x_372; x_301 = x_373; -x_302 = x_374; +x_302 = lean_box(0); x_303 = x_375; x_304 = x_376; x_305 = x_377; -x_306 = x_379; -x_307 = x_378; +x_306 = x_378; +x_307 = x_379; x_308 = x_380; x_309 = x_381; x_310 = x_382; @@ -4640,53 +4640,53 @@ goto block_359; else { lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; -x_389 = lean_ctor_get(x_362, 0); +x_389 = lean_ctor_get(x_360, 0); lean_inc(x_389); -lean_dec_ref(x_362); +lean_dec_ref(x_360); x_390 = l_Lean_Elab_Command_elabMacro___closed__56; x_391 = l_Lean_Elab_Command_elabMacro___closed__48; -lean_inc(x_369); +lean_inc(x_379); x_392 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_392, 0, x_369); +lean_ctor_set(x_392, 0, x_379); lean_ctor_set(x_392, 1, x_391); x_393 = l_Lean_Elab_Command_elabMacro___closed__57; -lean_inc(x_369); +lean_inc(x_379); x_394 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_394, 0, x_369); +lean_ctor_set(x_394, 0, x_379); lean_ctor_set(x_394, 1, x_393); x_395 = l_Lean_Elab_Command_elabMacro___closed__50; -lean_inc(x_369); +lean_inc(x_379); x_396 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_396, 0, x_369); +lean_ctor_set(x_396, 0, x_379); lean_ctor_set(x_396, 1, x_395); x_397 = l_Lean_Elab_Command_elabMacro___closed__51; -lean_inc(x_369); +lean_inc(x_379); x_398 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_398, 0, x_369); +lean_ctor_set(x_398, 0, x_379); lean_ctor_set(x_398, 1, x_397); -lean_inc(x_369); -x_399 = l_Lean_Syntax_node5(x_369, x_390, x_392, x_394, x_396, x_389, x_398); +lean_inc(x_379); +x_399 = l_Lean_Syntax_node5(x_379, x_390, x_392, x_394, x_396, x_389, x_398); x_400 = l_Array_mkArray1___redArg(x_399); -x_288 = x_360; -x_289 = x_361; -x_290 = x_364; -x_291 = x_363; +x_288 = x_361; +x_289 = x_362; +x_290 = x_363; +x_291 = x_364; x_292 = x_365; -x_293 = x_387; -x_294 = x_366; +x_293 = x_366; +x_294 = x_367; x_295 = x_368; -x_296 = x_367; +x_296 = x_387; x_297 = x_369; -x_298 = lean_box(0); +x_298 = x_370; x_299 = x_371; x_300 = x_372; x_301 = x_373; -x_302 = x_374; +x_302 = lean_box(0); x_303 = x_375; x_304 = x_376; x_305 = x_377; -x_306 = x_379; -x_307 = x_378; +x_306 = x_378; +x_307 = x_379; x_308 = x_380; x_309 = x_381; x_310 = x_382; @@ -4699,91 +4699,91 @@ goto block_359; block_438: { lean_object* x_428; lean_object* x_429; lean_object* x_430; -lean_inc_ref(x_424); -x_428 = l_Array_append___redArg(x_424, x_427); +lean_inc_ref(x_406); +x_428 = l_Array_append___redArg(x_406, x_427); lean_dec_ref(x_427); -lean_inc(x_419); -lean_inc(x_409); +lean_inc(x_422); +lean_inc(x_423); x_429 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_429, 0, x_409); -lean_ctor_set(x_429, 1, x_419); +lean_ctor_set(x_429, 0, x_423); +lean_ctor_set(x_429, 1, x_422); lean_ctor_set(x_429, 2, x_428); -lean_inc(x_409); +lean_inc(x_423); x_430 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_430, 0, x_409); -lean_ctor_set(x_430, 1, x_423); -if (lean_obj_tag(x_411) == 0) +lean_ctor_set(x_430, 0, x_423); +lean_ctor_set(x_430, 1, x_408); +if (lean_obj_tag(x_404) == 0) { lean_object* x_431; x_431 = l_Lean_Elab_Command_elabMacro___closed__11; x_360 = x_402; x_361 = x_403; -x_362 = x_404; -x_363 = x_429; -x_364 = x_405; -x_365 = x_406; -x_366 = x_407; -x_367 = x_408; -x_368 = x_430; -x_369 = x_409; -x_370 = lean_box(0); -x_371 = x_412; -x_372 = x_413; -x_373 = x_414; -x_374 = x_415; -x_375 = x_416; -x_376 = x_417; -x_377 = x_418; -x_378 = x_419; -x_379 = x_420; -x_380 = x_421; -x_381 = x_422; -x_382 = x_424; -x_383 = x_425; -x_384 = x_426; +x_362 = x_405; +x_363 = x_406; +x_364 = x_407; +x_365 = x_409; +x_366 = x_410; +x_367 = x_411; +x_368 = x_412; +x_369 = x_413; +x_370 = x_414; +x_371 = x_415; +x_372 = x_416; +x_373 = x_417; +x_374 = lean_box(0); +x_375 = x_419; +x_376 = x_420; +x_377 = x_421; +x_378 = x_422; +x_379 = x_423; +x_380 = x_430; +x_381 = x_424; +x_382 = x_425; +x_383 = x_426; +x_384 = x_429; x_385 = x_431; goto block_401; } else { lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; -x_432 = lean_ctor_get(x_411, 0); +x_432 = lean_ctor_get(x_404, 0); lean_inc(x_432); -lean_dec_ref(x_411); +lean_dec_ref(x_404); x_433 = l_Lean_Elab_Command_elabMacro___closed__59; x_434 = l_Lean_Elab_Command_elabMacro___closed__52; -lean_inc(x_409); +lean_inc(x_423); x_435 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_435, 0, x_409); +lean_ctor_set(x_435, 0, x_423); lean_ctor_set(x_435, 1, x_434); -lean_inc(x_409); -x_436 = l_Lean_Syntax_node2(x_409, x_433, x_435, x_432); +lean_inc(x_423); +x_436 = l_Lean_Syntax_node2(x_423, x_433, x_435, x_432); x_437 = l_Array_mkArray1___redArg(x_436); x_360 = x_402; x_361 = x_403; -x_362 = x_404; -x_363 = x_429; -x_364 = x_405; -x_365 = x_406; -x_366 = x_407; -x_367 = x_408; -x_368 = x_430; -x_369 = x_409; -x_370 = lean_box(0); -x_371 = x_412; -x_372 = x_413; -x_373 = x_414; -x_374 = x_415; -x_375 = x_416; -x_376 = x_417; -x_377 = x_418; -x_378 = x_419; -x_379 = x_420; -x_380 = x_421; -x_381 = x_422; -x_382 = x_424; -x_383 = x_425; -x_384 = x_426; +x_362 = x_405; +x_363 = x_406; +x_364 = x_407; +x_365 = x_409; +x_366 = x_410; +x_367 = x_411; +x_368 = x_412; +x_369 = x_413; +x_370 = x_414; +x_371 = x_415; +x_372 = x_416; +x_373 = x_417; +x_374 = lean_box(0); +x_375 = x_419; +x_376 = x_420; +x_377 = x_421; +x_378 = x_422; +x_379 = x_423; +x_380 = x_430; +x_381 = x_424; +x_382 = x_425; +x_383 = x_426; +x_384 = x_429; x_385 = x_437; goto block_401; } @@ -4791,36 +4791,36 @@ goto block_401; block_479: { lean_object* x_465; lean_object* x_466; -lean_inc_ref(x_461); -x_465 = l_Array_append___redArg(x_461, x_464); +lean_inc_ref(x_443); +x_465 = l_Array_append___redArg(x_443, x_464); lean_dec_ref(x_464); -lean_inc(x_456); -lean_inc(x_445); +lean_inc(x_459); +lean_inc(x_460); x_466 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_466, 0, x_445); -lean_ctor_set(x_466, 1, x_456); +lean_ctor_set(x_466, 0, x_460); +lean_ctor_set(x_466, 1, x_459); lean_ctor_set(x_466, 2, x_465); -if (lean_obj_tag(x_448) == 0) +if (lean_obj_tag(x_452) == 0) { lean_object* x_467; x_467 = l_Lean_Elab_Command_elabMacro___closed__11; x_402 = x_439; -x_403 = x_440; -x_404 = x_441; +x_403 = x_441; +x_404 = x_440; x_405 = x_442; x_406 = x_443; x_407 = x_444; -x_408 = x_466; -x_409 = x_445; -x_410 = lean_box(0); -x_411 = x_447; -x_412 = x_450; -x_413 = x_449; -x_414 = x_451; -x_415 = x_452; +x_408 = x_445; +x_409 = x_446; +x_410 = x_447; +x_411 = x_448; +x_412 = x_449; +x_413 = x_450; +x_414 = x_466; +x_415 = x_451; x_416 = x_453; x_417 = x_454; -x_418 = x_455; +x_418 = lean_box(0); x_419 = x_456; x_420 = x_457; x_421 = x_458; @@ -4835,51 +4835,51 @@ goto block_438; 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; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; -x_468 = lean_ctor_get(x_448, 0); +x_468 = lean_ctor_get(x_452, 0); lean_inc(x_468); -lean_dec_ref(x_448); +lean_dec_ref(x_452); x_469 = l_Lean_Elab_Command_elabMacro___closed__60; -lean_inc_ref(x_455); -x_470 = l_Lean_Name_mkStr4(x_5, x_6, x_455, x_469); +lean_inc_ref(x_446); +x_470 = l_Lean_Name_mkStr4(x_5, x_6, x_446, x_469); x_471 = l_Lean_Elab_Command_elabMacro___closed__61; -lean_inc(x_445); +lean_inc(x_460); x_472 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_472, 0, x_445); +lean_ctor_set(x_472, 0, x_460); lean_ctor_set(x_472, 1, x_471); -lean_inc_ref(x_461); -x_473 = l_Array_append___redArg(x_461, x_468); +lean_inc_ref(x_443); +x_473 = l_Array_append___redArg(x_443, x_468); lean_dec(x_468); -lean_inc(x_456); -lean_inc(x_445); +lean_inc(x_459); +lean_inc(x_460); x_474 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_474, 0, x_445); -lean_ctor_set(x_474, 1, x_456); +lean_ctor_set(x_474, 0, x_460); +lean_ctor_set(x_474, 1, x_459); lean_ctor_set(x_474, 2, x_473); x_475 = l_Lean_Elab_Command_elabMacro___closed__62; -lean_inc(x_445); +lean_inc(x_460); x_476 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_476, 0, x_445); +lean_ctor_set(x_476, 0, x_460); lean_ctor_set(x_476, 1, x_475); -lean_inc(x_445); -x_477 = l_Lean_Syntax_node3(x_445, x_470, x_472, x_474, x_476); +lean_inc(x_460); +x_477 = l_Lean_Syntax_node3(x_460, x_470, x_472, x_474, x_476); x_478 = l_Array_mkArray1___redArg(x_477); x_402 = x_439; -x_403 = x_440; -x_404 = x_441; +x_403 = x_441; +x_404 = x_440; x_405 = x_442; x_406 = x_443; x_407 = x_444; -x_408 = x_466; -x_409 = x_445; -x_410 = lean_box(0); -x_411 = x_447; -x_412 = x_450; -x_413 = x_449; -x_414 = x_451; -x_415 = x_452; +x_408 = x_445; +x_409 = x_446; +x_410 = x_447; +x_411 = x_448; +x_412 = x_449; +x_413 = x_450; +x_414 = x_466; +x_415 = x_451; x_416 = x_453; x_417 = x_454; -x_418 = x_455; +x_418 = lean_box(0); x_419 = x_456; x_420 = x_457; x_421 = x_458; @@ -4898,69 +4898,69 @@ lean_object* x_502; lean_object* x_503; lean_object* x_504; x_502 = l_Lean_Elab_Command_elabMacro___closed__63; x_503 = l_Lean_Elab_Command_elabMacro___closed__64; x_504 = l_Lean_Elab_Command_elabMacro___closed__65; -if (lean_obj_tag(x_488) == 0) +if (lean_obj_tag(x_491) == 0) { lean_object* x_505; x_505 = l_Lean_Elab_Command_elabMacro___closed__11; -x_439 = x_490; -x_440 = x_491; -x_441 = x_492; -x_442 = x_495; -x_443 = x_497; -x_444 = x_499; -x_445 = x_480; -x_446 = lean_box(0); -x_447 = x_481; -x_448 = x_482; -x_449 = x_483; -x_450 = x_484; -x_451 = x_485; -x_452 = x_486; -x_453 = x_487; -x_454 = x_488; -x_455 = x_489; -x_456 = x_494; -x_457 = x_493; -x_458 = x_503; -x_459 = x_496; -x_460 = x_502; -x_461 = x_504; -x_462 = x_498; -x_463 = x_500; +x_439 = x_480; +x_440 = x_481; +x_441 = x_482; +x_442 = x_483; +x_443 = x_504; +x_444 = x_486; +x_445 = x_502; +x_446 = x_489; +x_447 = x_491; +x_448 = x_493; +x_449 = x_494; +x_450 = x_495; +x_451 = x_499; +x_452 = x_500; +x_453 = x_484; +x_454 = x_485; +x_455 = lean_box(0); +x_456 = x_488; +x_457 = x_503; +x_458 = x_487; +x_459 = x_490; +x_460 = x_492; +x_461 = x_496; +x_462 = x_497; +x_463 = x_498; x_464 = x_505; goto block_479; } else { lean_object* x_506; lean_object* x_507; -x_506 = lean_ctor_get(x_488, 0); +x_506 = lean_ctor_get(x_491, 0); lean_inc(x_506); x_507 = l_Array_mkArray1___redArg(x_506); -x_439 = x_490; -x_440 = x_491; -x_441 = x_492; -x_442 = x_495; -x_443 = x_497; -x_444 = x_499; -x_445 = x_480; -x_446 = lean_box(0); -x_447 = x_481; -x_448 = x_482; -x_449 = x_483; -x_450 = x_484; -x_451 = x_485; -x_452 = x_486; -x_453 = x_487; -x_454 = x_488; -x_455 = x_489; -x_456 = x_494; -x_457 = x_493; -x_458 = x_503; -x_459 = x_496; -x_460 = x_502; -x_461 = x_504; -x_462 = x_498; -x_463 = x_500; +x_439 = x_480; +x_440 = x_481; +x_441 = x_482; +x_442 = x_483; +x_443 = x_504; +x_444 = x_486; +x_445 = x_502; +x_446 = x_489; +x_447 = x_491; +x_448 = x_493; +x_449 = x_494; +x_450 = x_495; +x_451 = x_499; +x_452 = x_500; +x_453 = x_484; +x_454 = x_485; +x_455 = lean_box(0); +x_456 = x_488; +x_457 = x_503; +x_458 = x_487; +x_459 = x_490; +x_460 = x_492; +x_461 = x_496; +x_462 = x_497; +x_463 = x_498; x_464 = x_507; goto block_479; } @@ -4980,13 +4980,13 @@ lean_dec(x_526); lean_dec(x_523); lean_dec_ref(x_522); lean_dec(x_521); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_515); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); lean_dec(x_1); x_529 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacro_spec__0___redArg(); @@ -5009,9 +5009,9 @@ lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; x_533 = lean_ctor_get(x_522, 5); x_534 = lean_ctor_get(x_522, 7); lean_dec(x_534); -x_535 = l_Lean_Syntax_getArg(x_526, x_510); -x_536 = lean_mk_empty_array_with_capacity(x_516); -x_537 = lean_array_push(x_536, x_515); +x_535 = l_Lean_Syntax_getArg(x_526, x_519); +x_536 = lean_mk_empty_array_with_capacity(x_513); +x_537 = lean_array_push(x_536, x_512); lean_inc(x_535); x_538 = lean_array_push(x_537, x_535); x_539 = l_Lean_Elab_Command_elabMacro___closed__69; @@ -5069,7 +5069,7 @@ if (lean_obj_tag(x_558) == 0) { lean_object* x_559; uint8_t x_560; lean_object* x_561; lean_dec_ref(x_558); -x_559 = l_Lean_Syntax_getArg(x_526, x_513); +x_559 = l_Lean_Syntax_getArg(x_526, x_511); lean_dec(x_526); x_560 = 0; x_561 = l_Lean_SourceInfo_fromRef(x_557, x_560); @@ -5079,53 +5079,53 @@ if (lean_obj_tag(x_533) == 0) lean_object* x_562; x_562 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabMacro_spec__1_spec__18___redArg(x_523); lean_dec_ref(x_562); -x_480 = x_561; -x_481 = x_512; -x_482 = x_511; -x_483 = x_513; -x_484 = x_554; -x_485 = x_535; -x_486 = x_555; -x_487 = x_550; -x_488 = x_518; -x_489 = x_519; -x_490 = x_520; -x_491 = x_522; -x_492 = x_509; -x_493 = x_540; -x_494 = x_539; -x_495 = x_514; -x_496 = x_545; -x_497 = x_559; -x_498 = x_533; -x_499 = x_517; -x_500 = x_523; +x_480 = x_509; +x_481 = x_510; +x_482 = x_540; +x_483 = x_511; +x_484 = x_514; +x_485 = x_522; +x_486 = x_523; +x_487 = x_516; +x_488 = x_550; +x_489 = x_517; +x_490 = x_539; +x_491 = x_518; +x_492 = x_561; +x_493 = x_545; +x_494 = x_555; +x_495 = x_533; +x_496 = x_554; +x_497 = x_535; +x_498 = x_559; +x_499 = x_515; +x_500 = x_520; x_501 = lean_box(0); goto block_508; } else { -x_480 = x_561; -x_481 = x_512; -x_482 = x_511; -x_483 = x_513; -x_484 = x_554; -x_485 = x_535; -x_486 = x_555; -x_487 = x_550; -x_488 = x_518; -x_489 = x_519; -x_490 = x_520; -x_491 = x_522; -x_492 = x_509; -x_493 = x_540; -x_494 = x_539; -x_495 = x_514; -x_496 = x_545; -x_497 = x_559; -x_498 = x_533; -x_499 = x_517; -x_500 = x_523; +x_480 = x_509; +x_481 = x_510; +x_482 = x_540; +x_483 = x_511; +x_484 = x_514; +x_485 = x_522; +x_486 = x_523; +x_487 = x_516; +x_488 = x_550; +x_489 = x_517; +x_490 = x_539; +x_491 = x_518; +x_492 = x_561; +x_493 = x_545; +x_494 = x_555; +x_495 = x_533; +x_496 = x_554; +x_497 = x_535; +x_498 = x_559; +x_499 = x_515; +x_500 = x_520; x_501 = lean_box(0); goto block_508; } @@ -5142,12 +5142,12 @@ lean_dec(x_535); lean_dec(x_533); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_563 = !lean_is_exclusive(x_558); if (x_563 == 0) @@ -5177,12 +5177,12 @@ lean_dec(x_535); lean_dec(x_533); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_566 = !lean_is_exclusive(x_556); if (x_566 == 0) @@ -5210,12 +5210,12 @@ lean_dec(x_535); lean_dec(x_533); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_569 = !lean_is_exclusive(x_551); if (x_569 == 0) @@ -5242,12 +5242,12 @@ lean_dec(x_535); lean_dec(x_533); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); lean_dec(x_1); x_572 = !lean_is_exclusive(x_544); @@ -5290,9 +5290,9 @@ lean_inc(x_577); lean_inc(x_576); lean_inc(x_575); lean_dec(x_522); -x_585 = l_Lean_Syntax_getArg(x_526, x_510); -x_586 = lean_mk_empty_array_with_capacity(x_516); -x_587 = lean_array_push(x_586, x_515); +x_585 = l_Lean_Syntax_getArg(x_526, x_519); +x_586 = lean_mk_empty_array_with_capacity(x_513); +x_587 = lean_array_push(x_586, x_512); lean_inc(x_585); x_588 = lean_array_push(x_587, x_585); x_589 = l_Lean_Elab_Command_elabMacro___closed__69; @@ -5361,7 +5361,7 @@ if (lean_obj_tag(x_609) == 0) { lean_object* x_610; uint8_t x_611; lean_object* x_612; lean_dec_ref(x_609); -x_610 = l_Lean_Syntax_getArg(x_526, x_513); +x_610 = l_Lean_Syntax_getArg(x_526, x_511); lean_dec(x_526); x_611 = 0; x_612 = l_Lean_SourceInfo_fromRef(x_608, x_611); @@ -5371,53 +5371,53 @@ if (lean_obj_tag(x_580) == 0) lean_object* x_613; x_613 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabMacro_spec__1_spec__18___redArg(x_523); lean_dec_ref(x_613); -x_480 = x_612; -x_481 = x_512; -x_482 = x_511; -x_483 = x_513; -x_484 = x_605; -x_485 = x_585; -x_486 = x_606; -x_487 = x_601; -x_488 = x_518; -x_489 = x_519; -x_490 = x_520; -x_491 = x_594; -x_492 = x_509; -x_493 = x_590; -x_494 = x_589; -x_495 = x_514; -x_496 = x_596; -x_497 = x_610; -x_498 = x_580; -x_499 = x_517; -x_500 = x_523; +x_480 = x_509; +x_481 = x_510; +x_482 = x_590; +x_483 = x_511; +x_484 = x_514; +x_485 = x_594; +x_486 = x_523; +x_487 = x_516; +x_488 = x_601; +x_489 = x_517; +x_490 = x_589; +x_491 = x_518; +x_492 = x_612; +x_493 = x_596; +x_494 = x_606; +x_495 = x_580; +x_496 = x_605; +x_497 = x_585; +x_498 = x_610; +x_499 = x_515; +x_500 = x_520; x_501 = lean_box(0); goto block_508; } else { -x_480 = x_612; -x_481 = x_512; -x_482 = x_511; -x_483 = x_513; -x_484 = x_605; -x_485 = x_585; -x_486 = x_606; -x_487 = x_601; -x_488 = x_518; -x_489 = x_519; -x_490 = x_520; -x_491 = x_594; -x_492 = x_509; -x_493 = x_590; -x_494 = x_589; -x_495 = x_514; -x_496 = x_596; -x_497 = x_610; -x_498 = x_580; -x_499 = x_517; -x_500 = x_523; +x_480 = x_509; +x_481 = x_510; +x_482 = x_590; +x_483 = x_511; +x_484 = x_514; +x_485 = x_594; +x_486 = x_523; +x_487 = x_516; +x_488 = x_601; +x_489 = x_517; +x_490 = x_589; +x_491 = x_518; +x_492 = x_612; +x_493 = x_596; +x_494 = x_606; +x_495 = x_580; +x_496 = x_605; +x_497 = x_585; +x_498 = x_610; +x_499 = x_515; +x_500 = x_520; x_501 = lean_box(0); goto block_508; } @@ -5434,12 +5434,12 @@ lean_dec(x_585); lean_dec(x_580); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_614 = lean_ctor_get(x_609, 0); lean_inc(x_614); @@ -5470,12 +5470,12 @@ lean_dec(x_585); lean_dec(x_580); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_617 = lean_ctor_get(x_607, 0); lean_inc(x_617); @@ -5504,12 +5504,12 @@ lean_dec(x_585); lean_dec(x_580); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); x_620 = lean_ctor_get(x_602, 0); lean_inc(x_620); @@ -5537,12 +5537,12 @@ lean_dec(x_585); lean_dec(x_580); lean_dec(x_526); lean_dec(x_523); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); -lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); lean_dec(x_1); x_623 = lean_ctor_get(x_595, 0); @@ -5571,13 +5571,13 @@ lean_dec(x_526); lean_dec(x_523); lean_dec_ref(x_522); lean_dec(x_521); -lean_dec_ref(x_519); +lean_dec(x_520); lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_515); +lean_dec_ref(x_517); +lean_dec(x_516); lean_dec(x_514); lean_dec(x_512); -lean_dec(x_511); +lean_dec(x_510); lean_dec(x_509); lean_dec(x_1); x_626 = !lean_is_exclusive(x_530); @@ -5608,7 +5608,7 @@ if (x_647 == 0) { uint8_t x_648; lean_inc(x_646); -x_648 = l_Lean_Syntax_matchesNull(x_646, x_633); +x_648 = l_Lean_Syntax_matchesNull(x_646, x_631); if (x_648 == 0) { lean_object* x_649; @@ -5616,13 +5616,13 @@ lean_dec(x_646); lean_dec(x_643); lean_dec_ref(x_642); lean_dec(x_641); -lean_dec_ref(x_639); +lean_dec(x_640); lean_dec(x_638); -lean_dec(x_637); +lean_dec_ref(x_637); lean_dec(x_635); lean_dec(x_634); -lean_dec(x_632); -lean_dec(x_631); +lean_dec(x_633); +lean_dec(x_630); lean_dec(x_1); x_649 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacro_spec__0___redArg(); return x_649; @@ -5642,13 +5642,13 @@ lean_dec(x_650); lean_dec(x_643); lean_dec_ref(x_642); lean_dec(x_641); -lean_dec_ref(x_639); +lean_dec(x_640); lean_dec(x_638); -lean_dec(x_637); +lean_dec_ref(x_637); lean_dec(x_635); lean_dec(x_634); -lean_dec(x_632); -lean_dec(x_631); +lean_dec(x_633); +lean_dec(x_630); lean_dec(x_1); x_653 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacro_spec__0___redArg(); return x_653; @@ -5656,20 +5656,20 @@ return x_653; else { lean_object* x_654; lean_object* x_655; -x_654 = l_Lean_Syntax_getArg(x_650, x_630); +x_654 = l_Lean_Syntax_getArg(x_650, x_639); lean_dec(x_650); x_655 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_655, 0, x_654); x_509 = x_641; x_510 = x_630; -x_511 = x_632; -x_512 = x_631; -x_513 = x_633; -x_514 = x_635; -x_515 = x_634; -x_516 = x_636; -x_517 = x_638; -x_518 = x_637; +x_511 = x_631; +x_512 = x_633; +x_513 = x_632; +x_514 = x_634; +x_515 = x_636; +x_516 = x_635; +x_517 = x_637; +x_518 = x_638; x_519 = x_639; x_520 = x_640; x_521 = x_655; @@ -5687,14 +5687,14 @@ lean_dec(x_646); x_656 = lean_box(0); x_509 = x_641; x_510 = x_630; -x_511 = x_632; -x_512 = x_631; -x_513 = x_633; -x_514 = x_635; -x_515 = x_634; -x_516 = x_636; -x_517 = x_638; -x_518 = x_637; +x_511 = x_631; +x_512 = x_633; +x_513 = x_632; +x_514 = x_634; +x_515 = x_636; +x_516 = x_635; +x_517 = x_637; +x_518 = x_638; x_519 = x_639; x_520 = x_640; x_521 = x_656; @@ -5714,7 +5714,7 @@ if (x_674 == 0) { uint8_t x_675; lean_inc(x_673); -x_675 = l_Lean_Syntax_matchesNull(x_673, x_660); +x_675 = l_Lean_Syntax_matchesNull(x_673, x_658); if (x_675 == 0) { lean_object* x_676; @@ -5722,10 +5722,10 @@ lean_dec(x_673); lean_dec(x_670); lean_dec_ref(x_669); lean_dec(x_668); -lean_dec_ref(x_666); +lean_dec(x_667); lean_dec(x_665); -lean_dec(x_664); -lean_dec(x_662); +lean_dec_ref(x_664); +lean_dec(x_663); lean_dec(x_661); lean_dec(x_659); lean_dec(x_1); @@ -5747,10 +5747,10 @@ lean_dec(x_677); lean_dec(x_670); lean_dec_ref(x_669); lean_dec(x_668); -lean_dec_ref(x_666); +lean_dec(x_667); lean_dec(x_665); -lean_dec(x_664); -lean_dec(x_662); +lean_dec_ref(x_664); +lean_dec(x_663); lean_dec(x_661); lean_dec(x_659); lean_dec(x_1); @@ -5760,19 +5760,19 @@ return x_680; else { lean_object* x_681; lean_object* x_682; -x_681 = l_Lean_Syntax_getArg(x_677, x_658); +x_681 = l_Lean_Syntax_getArg(x_677, x_666); lean_dec(x_677); x_682 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_682, 0, x_681); -x_630 = x_658; -x_631 = x_668; -x_632 = x_659; -x_633 = x_660; -x_634 = x_662; -x_635 = x_661; -x_636 = x_663; -x_637 = x_665; -x_638 = x_664; +x_630 = x_668; +x_631 = x_658; +x_632 = x_660; +x_633 = x_659; +x_634 = x_661; +x_635 = x_663; +x_636 = x_662; +x_637 = x_664; +x_638 = x_665; x_639 = x_666; x_640 = x_667; x_641 = x_682; @@ -5788,15 +5788,15 @@ else lean_object* x_683; lean_dec(x_673); x_683 = lean_box(0); -x_630 = x_658; -x_631 = x_668; -x_632 = x_659; -x_633 = x_660; -x_634 = x_662; -x_635 = x_661; -x_636 = x_663; -x_637 = x_665; -x_638 = x_664; +x_630 = x_668; +x_631 = x_658; +x_632 = x_660; +x_633 = x_659; +x_634 = x_661; +x_635 = x_663; +x_636 = x_662; +x_637 = x_664; +x_638 = x_665; x_639 = x_666; x_640 = x_667; x_641 = x_683; @@ -5883,16 +5883,16 @@ x_708 = l_Lean_Syntax_getArg(x_704, x_685); lean_dec(x_704); x_709 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_709, 0, x_708); -x_658 = x_697; -x_659 = x_687; -x_660 = x_685; +x_658 = x_685; +x_659 = x_698; +x_660 = x_691; x_661 = x_692; -x_662 = x_698; -x_663 = x_691; -x_664 = x_694; +x_662 = x_695; +x_663 = x_694; +x_664 = x_693; x_665 = x_686; -x_666 = x_693; -x_667 = x_695; +x_666 = x_697; +x_667 = x_687; x_668 = x_709; x_669 = x_688; x_670 = x_689; @@ -5906,16 +5906,16 @@ else lean_object* x_710; lean_dec(x_700); x_710 = lean_box(0); -x_658 = x_697; -x_659 = x_687; -x_660 = x_685; +x_658 = x_685; +x_659 = x_698; +x_660 = x_691; x_661 = x_692; -x_662 = x_698; -x_663 = x_691; -x_664 = x_694; +x_662 = x_695; +x_663 = x_694; +x_664 = x_693; x_665 = x_686; -x_666 = x_693; -x_667 = x_695; +x_666 = x_697; +x_667 = x_687; x_668 = x_710; x_669 = x_688; x_670 = x_689; @@ -6002,74 +6002,74 @@ goto block_711; block_47: { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; 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_inc_ref(x_17); -x_21 = l_Array_append___redArg(x_17, x_20); +lean_inc_ref(x_10); +x_21 = l_Array_append___redArg(x_10, x_20); lean_dec_ref(x_20); -lean_inc(x_14); -lean_inc(x_7); +lean_inc(x_16); +lean_inc(x_8); x_22 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_22, 0, x_7); -lean_ctor_set(x_22, 1, x_14); +lean_ctor_set(x_22, 0, x_8); +lean_ctor_set(x_22, 1, x_16); lean_ctor_set(x_22, 2, x_21); -lean_inc(x_14); -lean_inc(x_7); +lean_inc(x_16); +lean_inc(x_8); x_23 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_23, 0, x_7); -lean_ctor_set(x_23, 1, x_14); -lean_ctor_set(x_23, 2, x_17); -lean_inc(x_7); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_16); +lean_ctor_set(x_23, 2, x_10); +lean_inc(x_8); x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_11); +lean_ctor_set(x_24, 0, x_8); +lean_ctor_set(x_24, 1, x_18); x_25 = l_Lean_Elab_Command_elabMacro___closed__2; -lean_inc_ref(x_9); -x_26 = l_Lean_Name_mkStr4(x_5, x_6, x_9, x_25); +lean_inc_ref(x_14); +x_26 = l_Lean_Name_mkStr4(x_5, x_6, x_14, x_25); x_27 = l_Lean_Elab_Command_elabMacro___closed__3; -lean_inc_ref(x_9); -x_28 = l_Lean_Name_mkStr4(x_5, x_6, x_9, x_27); +lean_inc_ref(x_14); +x_28 = l_Lean_Name_mkStr4(x_5, x_6, x_14, x_27); x_29 = l_Lean_Elab_Command_elabMacro___closed__4; -lean_inc(x_7); +lean_inc(x_8); x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_7); +lean_ctor_set(x_30, 0, x_8); lean_ctor_set(x_30, 1, x_29); x_31 = l_Lean_Elab_Command_elabMacro___closed__5; -x_32 = l_Lean_Name_mkStr4(x_5, x_6, x_9, x_31); +x_32 = l_Lean_Name_mkStr4(x_5, x_6, x_14, x_31); x_33 = l_Lean_Elab_Command_elabMacro___closed__6; -lean_inc(x_7); +lean_inc(x_8); x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_7); +lean_ctor_set(x_34, 0, x_8); lean_ctor_set(x_34, 1, x_33); -lean_inc(x_7); +lean_inc(x_8); x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_7); -lean_ctor_set(x_35, 1, x_18); +lean_ctor_set(x_35, 0, x_8); +lean_ctor_set(x_35, 1, x_7); lean_inc_ref(x_35); lean_inc_ref(x_34); lean_inc(x_32); -lean_inc(x_7); -x_36 = l_Lean_Syntax_node3(x_7, x_32, x_34, x_19, x_35); -lean_inc(x_14); -lean_inc(x_7); -x_37 = l_Lean_Syntax_node1(x_7, x_14, x_36); -lean_inc(x_14); -lean_inc(x_7); -x_38 = l_Lean_Syntax_node1(x_7, x_14, x_37); +lean_inc(x_8); +x_36 = l_Lean_Syntax_node3(x_8, x_32, x_34, x_17, x_35); +lean_inc(x_16); +lean_inc(x_8); +x_37 = l_Lean_Syntax_node1(x_8, x_16, x_36); +lean_inc(x_16); +lean_inc(x_8); +x_38 = l_Lean_Syntax_node1(x_8, x_16, x_37); x_39 = l_Lean_Elab_Command_elabMacro___closed__7; -lean_inc(x_7); +lean_inc(x_8); x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_7); +lean_ctor_set(x_40, 0, x_8); lean_ctor_set(x_40, 1, x_39); -lean_inc(x_7); -x_41 = l_Lean_Syntax_node3(x_7, x_32, x_34, x_13, x_35); -lean_inc(x_7); -x_42 = l_Lean_Syntax_node4(x_7, x_28, x_30, x_38, x_40, x_41); -lean_inc(x_7); -x_43 = l_Lean_Syntax_node1(x_7, x_14, x_42); -lean_inc(x_7); -x_44 = l_Lean_Syntax_node1(x_7, x_26, x_43); +lean_inc(x_8); +x_41 = l_Lean_Syntax_node3(x_8, x_32, x_34, x_9, x_35); +lean_inc(x_8); +x_42 = l_Lean_Syntax_node4(x_8, x_28, x_30, x_38, x_40, x_41); +lean_inc(x_8); +x_43 = l_Lean_Syntax_node1(x_8, x_16, x_42); +lean_inc(x_8); +x_44 = l_Lean_Syntax_node1(x_8, x_26, x_43); lean_inc_ref(x_23); -x_45 = l_Lean_Syntax_node6(x_7, x_16, x_22, x_23, x_15, x_24, x_23, x_44); -x_46 = l_Lean_Elab_Command_elabCommand(x_45, x_10, x_12); +x_45 = l_Lean_Syntax_node6(x_8, x_15, x_22, x_23, x_19, x_24, x_23, x_44); +x_46 = l_Lean_Elab_Command_elabCommand(x_45, x_12, x_13); return x_46; } block_65: @@ -6077,46 +6077,46 @@ return x_46; lean_object* x_60; lean_object* x_61; x_60 = l_Lean_Elab_Command_elabMacro___closed__9; x_61 = l_Lean_Elab_Command_elabMacro___closed__10; -if (lean_obj_tag(x_53) == 0) +if (lean_obj_tag(x_56) == 0) { lean_object* x_62; x_62 = l_Lean_Elab_Command_elabMacro___closed__11; -x_7 = x_49; -x_8 = lean_box(0); -x_9 = x_57; -x_10 = x_56; -x_11 = x_60; -x_12 = x_48; -x_13 = x_50; -x_14 = x_51; -x_15 = x_52; -x_16 = x_61; -x_17 = x_54; -x_18 = x_55; -x_19 = x_58; +x_7 = x_48; +x_8 = x_49; +x_9 = x_51; +x_10 = x_52; +x_11 = lean_box(0); +x_12 = x_54; +x_13 = x_53; +x_14 = x_55; +x_15 = x_61; +x_16 = x_58; +x_17 = x_50; +x_18 = x_60; +x_19 = x_57; x_20 = x_62; goto block_47; } else { lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_53, 0); +x_63 = lean_ctor_get(x_56, 0); lean_inc(x_63); -lean_dec_ref(x_53); +lean_dec_ref(x_56); x_64 = l_Array_mkArray1___redArg(x_63); -x_7 = x_49; -x_8 = lean_box(0); -x_9 = x_57; -x_10 = x_56; -x_11 = x_60; -x_12 = x_48; -x_13 = x_50; -x_14 = x_51; -x_15 = x_52; -x_16 = x_61; -x_17 = x_54; -x_18 = x_55; -x_19 = x_58; +x_7 = x_48; +x_8 = x_49; +x_9 = x_51; +x_10 = x_52; +x_11 = lean_box(0); +x_12 = x_54; +x_13 = x_53; +x_14 = x_55; +x_15 = x_61; +x_16 = x_58; +x_17 = x_50; +x_18 = x_60; +x_19 = x_57; x_20 = x_64; goto block_47; } @@ -6124,97 +6124,97 @@ goto block_47; block_154: { lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; 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_inc_ref(x_77); -x_84 = l_Array_append___redArg(x_77, x_83); +lean_inc_ref(x_69); +x_84 = l_Array_append___redArg(x_69, x_83); lean_dec_ref(x_83); lean_inc(x_73); -lean_inc(x_78); +lean_inc(x_74); x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_78); +lean_ctor_set(x_85, 0, x_74); lean_ctor_set(x_85, 1, x_73); lean_ctor_set(x_85, 2, x_84); lean_inc(x_73); -lean_inc(x_78); +lean_inc(x_74); x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_78); +lean_ctor_set(x_86, 0, x_74); lean_ctor_set(x_86, 1, x_73); -lean_ctor_set(x_86, 2, x_77); -lean_inc(x_78); +lean_ctor_set(x_86, 2, x_69); +lean_inc(x_74); x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_78); -lean_ctor_set(x_87, 1, x_66); +lean_ctor_set(x_87, 0, x_74); +lean_ctor_set(x_87, 1, x_80); x_88 = l_Lean_Elab_Command_elabMacro___closed__2; -lean_inc_ref(x_69); -x_89 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_88); +lean_inc_ref(x_72); +x_89 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_88); x_90 = l_Lean_Elab_Command_elabMacro___closed__3; -lean_inc_ref(x_69); -x_91 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_90); +lean_inc_ref(x_72); +x_91 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_90); x_92 = l_Lean_Elab_Command_elabMacro___closed__4; -lean_inc(x_78); +lean_inc(x_74); x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_78); +lean_ctor_set(x_93, 0, x_74); lean_ctor_set(x_93, 1, x_92); x_94 = l_Lean_Elab_Command_elabMacro___closed__5; -lean_inc_ref(x_69); -x_95 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_94); +lean_inc_ref(x_72); +x_95 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_94); x_96 = l_Lean_Elab_Command_elabMacro___closed__6; -lean_inc(x_78); +lean_inc(x_74); x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_78); +lean_ctor_set(x_97, 0, x_74); lean_ctor_set(x_97, 1, x_96); -lean_inc(x_78); +lean_inc(x_74); x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_78); -lean_ctor_set(x_98, 1, x_79); +lean_ctor_set(x_98, 0, x_74); +lean_ctor_set(x_98, 1, x_66); lean_inc_ref(x_98); -lean_inc(x_78); -x_99 = l_Lean_Syntax_node3(x_78, x_95, x_97, x_82, x_98); +lean_inc(x_74); +x_99 = l_Lean_Syntax_node3(x_74, x_95, x_97, x_77, x_98); lean_inc(x_73); -lean_inc(x_78); -x_100 = l_Lean_Syntax_node1(x_78, x_73, x_99); +lean_inc(x_74); +x_100 = l_Lean_Syntax_node1(x_74, x_73, x_99); lean_inc(x_73); -lean_inc(x_78); -x_101 = l_Lean_Syntax_node1(x_78, x_73, x_100); +lean_inc(x_74); +x_101 = l_Lean_Syntax_node1(x_74, x_73, x_100); x_102 = l_Lean_Elab_Command_elabMacro___closed__7; -lean_inc(x_78); +lean_inc(x_74); x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_78); +lean_ctor_set(x_103, 0, x_74); lean_ctor_set(x_103, 1, x_102); x_104 = l_Lean_Elab_Command_elabMacro___closed__12; -lean_inc_ref(x_69); -x_105 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_104); +lean_inc_ref(x_72); +x_105 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_104); x_106 = l_Lean_Elab_Command_elabMacro___closed__14; x_107 = l_Lean_Elab_Command_elabMacro___closed__17; -lean_inc(x_70); +lean_inc(x_75); lean_inc(x_81); -x_108 = l_Lean_addMacroScope(x_81, x_107, x_70); +x_108 = l_Lean_addMacroScope(x_81, x_107, x_75); x_109 = l_Lean_Elab_Command_elabMacro___closed__19; -lean_inc(x_78); +lean_inc(x_74); x_110 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_110, 0, x_78); +lean_ctor_set(x_110, 0, x_74); lean_ctor_set(x_110, 1, x_106); lean_ctor_set(x_110, 2, x_108); lean_ctor_set(x_110, 3, x_109); x_111 = l_Lean_Elab_Command_elabMacro___closed__20; -lean_inc_ref(x_69); -x_112 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_111); +lean_inc_ref(x_72); +x_112 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_111); x_113 = l_Lean_Elab_Command_elabMacro___closed__21; -lean_inc_ref(x_69); -x_114 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_113); -lean_inc(x_78); +lean_inc_ref(x_72); +x_114 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_113); +lean_inc(x_74); x_115 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_115, 0, x_78); -lean_ctor_set(x_115, 1, x_67); +lean_ctor_set(x_115, 0, x_74); +lean_ctor_set(x_115, 1, x_79); x_116 = l_Lean_Elab_Command_elabMacro___closed__23; x_117 = l_Lean_Elab_Command_elabMacro___closed__24; x_118 = lean_box(0); -lean_inc(x_70); +lean_inc(x_75); lean_inc(x_81); -x_119 = l_Lean_addMacroScope(x_81, x_118, x_70); +x_119 = l_Lean_addMacroScope(x_81, x_118, x_75); x_120 = l_Lean_Elab_Command_elabMacro___closed__27; x_121 = l_Lean_Elab_Command_elabMacro___closed__29; -lean_inc_ref(x_69); -x_122 = l_Lean_Name_mkStr3(x_5, x_6, x_69); +lean_inc_ref(x_72); +x_122 = l_Lean_Name_mkStr3(x_5, x_6, x_72); x_123 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_123, 0, x_122); x_124 = l_Lean_Elab_Command_elabMacro___closed__33; @@ -6227,61 +6227,61 @@ lean_ctor_set(x_126, 1, x_125); x_127 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_127, 0, x_120); lean_ctor_set(x_127, 1, x_126); -lean_inc(x_78); +lean_inc(x_74); x_128 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_128, 0, x_78); +lean_ctor_set(x_128, 0, x_74); lean_ctor_set(x_128, 1, x_117); lean_ctor_set(x_128, 2, x_119); lean_ctor_set(x_128, 3, x_127); -lean_inc(x_78); -x_129 = l_Lean_Syntax_node1(x_78, x_116, x_128); -lean_inc(x_78); -x_130 = l_Lean_Syntax_node2(x_78, x_114, x_115, x_129); +lean_inc(x_74); +x_129 = l_Lean_Syntax_node1(x_74, x_116, x_128); +lean_inc(x_74); +x_130 = l_Lean_Syntax_node2(x_74, x_114, x_115, x_129); x_131 = l_Lean_Elab_Command_elabMacro___closed__34; -x_132 = l_Lean_Name_mkStr4(x_5, x_6, x_69, x_131); +x_132 = l_Lean_Name_mkStr4(x_5, x_6, x_72, x_131); x_133 = l_Lean_Elab_Command_elabMacro___closed__35; -lean_inc(x_78); +lean_inc(x_74); x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_78); +lean_ctor_set(x_134, 0, x_74); lean_ctor_set(x_134, 1, x_133); x_135 = l_Lean_Elab_Command_elabMacro___closed__37; x_136 = l_Lean_Elab_Command_elabMacro___closed__40; -x_137 = l_Lean_addMacroScope(x_81, x_136, x_70); +x_137 = l_Lean_addMacroScope(x_81, x_136, x_75); x_138 = l_Lean_Elab_Command_elabMacro___closed__43; -lean_inc(x_78); +lean_inc(x_74); x_139 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_139, 0, x_78); +lean_ctor_set(x_139, 0, x_74); lean_ctor_set(x_139, 1, x_135); lean_ctor_set(x_139, 2, x_137); lean_ctor_set(x_139, 3, x_138); -lean_inc(x_78); -x_140 = l_Lean_Syntax_node2(x_78, x_132, x_134, x_139); -x_141 = l_Lean_TSyntax_getId(x_76); -lean_dec(x_76); +lean_inc(x_74); +x_140 = l_Lean_Syntax_node2(x_74, x_132, x_134, x_139); +x_141 = l_Lean_TSyntax_getId(x_78); +lean_dec(x_78); x_142 = lean_erase_macro_scopes(x_141); x_143 = l_Lean_instQuoteNameMkStr1___private__1(x_142); lean_inc(x_73); -lean_inc(x_78); -x_144 = l_Lean_Syntax_node1(x_78, x_73, x_143); +lean_inc(x_74); +x_144 = l_Lean_Syntax_node1(x_74, x_73, x_143); lean_inc(x_105); -lean_inc(x_78); -x_145 = l_Lean_Syntax_node2(x_78, x_105, x_140, x_144); -lean_inc(x_78); -x_146 = l_Lean_Syntax_node3(x_78, x_112, x_130, x_145, x_98); +lean_inc(x_74); +x_145 = l_Lean_Syntax_node2(x_74, x_105, x_140, x_144); +lean_inc(x_74); +x_146 = l_Lean_Syntax_node3(x_74, x_112, x_130, x_145, x_98); lean_inc(x_73); -lean_inc(x_78); -x_147 = l_Lean_Syntax_node2(x_78, x_73, x_146, x_80); -lean_inc(x_78); -x_148 = l_Lean_Syntax_node2(x_78, x_105, x_110, x_147); -lean_inc(x_78); -x_149 = l_Lean_Syntax_node4(x_78, x_91, x_93, x_101, x_103, x_148); -lean_inc(x_78); -x_150 = l_Lean_Syntax_node1(x_78, x_73, x_149); -lean_inc(x_78); -x_151 = l_Lean_Syntax_node1(x_78, x_89, x_150); +lean_inc(x_74); +x_147 = l_Lean_Syntax_node2(x_74, x_73, x_146, x_76); +lean_inc(x_74); +x_148 = l_Lean_Syntax_node2(x_74, x_105, x_110, x_147); +lean_inc(x_74); +x_149 = l_Lean_Syntax_node4(x_74, x_91, x_93, x_101, x_103, x_148); +lean_inc(x_74); +x_150 = l_Lean_Syntax_node1(x_74, x_73, x_149); +lean_inc(x_74); +x_151 = l_Lean_Syntax_node1(x_74, x_89, x_150); lean_inc_ref(x_86); -x_152 = l_Lean_Syntax_node6(x_78, x_72, x_85, x_86, x_74, x_87, x_86, x_151); -x_153 = l_Lean_Elab_Command_elabCommand(x_152, x_68, x_71); +x_152 = l_Lean_Syntax_node6(x_74, x_68, x_85, x_86, x_82, x_87, x_86, x_151); +x_153 = l_Lean_Elab_Command_elabCommand(x_152, x_71, x_70); return x_153; } block_176: @@ -6289,25 +6289,25 @@ return x_153; lean_object* x_171; lean_object* x_172; x_171 = l_Lean_Elab_Command_elabMacro___closed__9; x_172 = l_Lean_Elab_Command_elabMacro___closed__10; -if (lean_obj_tag(x_156) == 0) +if (lean_obj_tag(x_160) == 0) { lean_object* x_173; x_173 = l_Lean_Elab_Command_elabMacro___closed__11; -x_66 = x_171; -x_67 = x_155; -x_68 = x_157; -x_69 = x_158; -x_70 = x_159; -x_71 = x_160; -x_72 = x_172; +x_66 = x_155; +x_67 = lean_box(0); +x_68 = x_172; +x_69 = x_156; +x_70 = x_157; +x_71 = x_158; +x_72 = x_159; x_73 = x_161; x_74 = x_162; -x_75 = lean_box(0); -x_76 = x_163; -x_77 = x_164; -x_78 = x_165; -x_79 = x_166; -x_80 = x_167; +x_75 = x_163; +x_76 = x_164; +x_77 = x_165; +x_78 = x_166; +x_79 = x_167; +x_80 = x_171; x_81 = x_169; x_82 = x_168; x_83 = x_173; @@ -6316,25 +6316,25 @@ goto block_154; else { lean_object* x_174; lean_object* x_175; -x_174 = lean_ctor_get(x_156, 0); +x_174 = lean_ctor_get(x_160, 0); lean_inc(x_174); -lean_dec_ref(x_156); +lean_dec_ref(x_160); x_175 = l_Array_mkArray1___redArg(x_174); -x_66 = x_171; -x_67 = x_155; -x_68 = x_157; -x_69 = x_158; -x_70 = x_159; -x_71 = x_160; -x_72 = x_172; +x_66 = x_155; +x_67 = lean_box(0); +x_68 = x_172; +x_69 = x_156; +x_70 = x_157; +x_71 = x_158; +x_72 = x_159; x_73 = x_161; x_74 = x_162; -x_75 = lean_box(0); -x_76 = x_163; -x_77 = x_164; -x_78 = x_165; -x_79 = x_166; -x_80 = x_167; +x_75 = x_163; +x_76 = x_164; +x_77 = x_165; +x_78 = x_166; +x_79 = x_167; +x_80 = x_171; x_81 = x_169; x_82 = x_168; x_83 = x_175; diff --git a/stage0/stdlib/Lean/Elab/MacroRules.c b/stage0/stdlib/Lean/Elab/MacroRules.c index 03bbca3d962c..89c962e1417e 100644 --- a/stage0/stdlib/Lean/Elab/MacroRules.c +++ b/stage0/stdlib/Lean/Elab/MacroRules.c @@ -1354,9 +1354,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_63; -x_26 = x_62; -x_27 = lean_box(0); +x_25 = x_62; +x_26 = lean_box(0); +x_27 = x_63; goto block_34; } else @@ -1372,9 +1372,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_63; -x_26 = x_62; -x_27 = lean_box(0); +x_25 = x_62; +x_26 = lean_box(0); +x_27 = x_63; goto block_34; } else @@ -1538,7 +1538,7 @@ x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__priv 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_throwErrorAt___at___00Lean_Elab_Command_elabMacroRulesAux_spec__1___redArg(x_10, x_32, x_26, x_25); +x_33 = l_Lean_throwErrorAt___at___00Lean_Elab_Command_elabMacroRulesAux_spec__1___redArg(x_10, x_32, x_25, x_27); lean_dec(x_10); x_20 = x_33; goto block_24; @@ -1747,9 +1747,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_62; -x_26 = x_63; -x_27 = lean_box(0); +x_25 = x_63; +x_26 = lean_box(0); +x_27 = x_62; goto block_34; } else @@ -1765,9 +1765,9 @@ lean_dec_ref(x_59); lean_dec(x_44); lean_dec_ref(x_12); lean_dec_ref(x_5); -x_25 = x_62; -x_26 = x_63; -x_27 = lean_box(0); +x_25 = x_63; +x_26 = lean_box(0); +x_27 = x_62; goto block_34; } else @@ -1930,7 +1930,7 @@ x_31 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__priv 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_throwErrorAt___at___00Lean_Elab_Command_elabMacroRulesAux_spec__1___redArg(x_10, x_32, x_25, x_26); +x_33 = l_Lean_throwErrorAt___at___00Lean_Elab_Command_elabMacroRulesAux_spec__1___redArg(x_10, x_32, x_27, x_25); lean_dec(x_10); x_20 = x_33; goto block_24; @@ -2393,9 +2393,9 @@ x_154 = lean_ctor_get(x_153, 0); lean_inc(x_154); lean_dec_ref(x_153); x_125 = x_152; -x_126 = x_146; -x_127 = x_150; -x_128 = x_151; +x_126 = x_150; +x_127 = x_151; +x_128 = x_146; x_129 = x_154; x_130 = lean_box(0); goto block_139; @@ -2407,9 +2407,9 @@ x_155 = lean_ctor_get(x_143, 0); lean_inc(x_155); lean_dec_ref(x_143); x_125 = x_152; -x_126 = x_146; -x_127 = x_150; -x_128 = x_151; +x_126 = x_150; +x_127 = x_151; +x_128 = x_146; x_129 = x_155; x_130 = lean_box(0); goto block_139; @@ -2539,78 +2539,78 @@ return x_140; block_124: { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_15); -x_27 = l_Array_append___redArg(x_15, x_26); +lean_inc_ref(x_25); +x_27 = l_Array_append___redArg(x_25, x_26); lean_dec_ref(x_26); -lean_inc(x_25); -lean_inc(x_17); +lean_inc(x_23); +lean_inc(x_16); x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_25); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_23); lean_ctor_set(x_28, 2, x_27); x_29 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__5; x_30 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__6; x_31 = l_Lean_Elab_Command_elabMacroRulesAux___closed__0; -lean_inc_ref(x_23); -x_32 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_31); +lean_inc_ref(x_20); +x_32 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_31); x_33 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; -lean_inc(x_17); +lean_inc(x_16); x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_17); +lean_ctor_set(x_34, 0, x_16); lean_ctor_set(x_34, 1, x_33); -lean_inc_ref(x_15); -x_35 = l_Array_append___redArg(x_15, x_20); -lean_dec_ref(x_20); -lean_inc(x_25); -lean_inc(x_17); +lean_inc_ref(x_25); +x_35 = l_Array_append___redArg(x_25, x_22); +lean_dec_ref(x_22); +lean_inc(x_23); +lean_inc(x_16); x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_17); -lean_ctor_set(x_36, 1, x_25); +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_36, 1, x_23); lean_ctor_set(x_36, 2, x_35); x_37 = l_Lean_Elab_Command_elabMacroRulesAux___closed__2; -lean_inc(x_17); +lean_inc(x_16); x_38 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_38, 0, x_17); +lean_ctor_set(x_38, 0, x_16); lean_ctor_set(x_38, 1, x_37); -lean_inc(x_17); -x_39 = l_Lean_Syntax_node3(x_17, x_32, x_34, x_36, x_38); -lean_inc(x_25); -lean_inc(x_17); -x_40 = l_Lean_Syntax_node1(x_17, x_25, x_39); -lean_inc(x_17); +lean_inc(x_16); +x_39 = l_Lean_Syntax_node3(x_16, x_32, x_34, x_36, x_38); +lean_inc(x_23); +lean_inc(x_16); +x_40 = l_Lean_Syntax_node1(x_16, x_23, x_39); +lean_inc(x_16); x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_17); -lean_ctor_set(x_41, 1, x_16); +lean_ctor_set(x_41, 0, x_16); +lean_ctor_set(x_41, 1, x_19); x_42 = l_Lean_Elab_Command_elabMacroRulesAux___closed__4; x_43 = l_Lean_Elab_Command_elabMacroRulesAux___closed__5; -lean_inc(x_22); -lean_inc(x_19); -x_44 = l_Lean_addMacroScope(x_19, x_43, x_22); +lean_inc(x_18); +lean_inc(x_21); +x_44 = l_Lean_addMacroScope(x_21, x_43, x_18); x_45 = lean_box(0); -lean_inc(x_17); +lean_inc(x_16); x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_17); +lean_ctor_set(x_46, 0, x_16); lean_ctor_set(x_46, 1, x_42); lean_ctor_set(x_46, 2, x_44); lean_ctor_set(x_46, 3, x_45); x_47 = 1; x_48 = l_Lean_mkIdentFrom(x_4, x_5, x_47); -lean_inc(x_25); -lean_inc(x_17); -x_49 = l_Lean_Syntax_node2(x_17, x_25, x_46, x_48); +lean_inc(x_23); +lean_inc(x_16); +x_49 = l_Lean_Syntax_node2(x_16, x_23, x_46, x_48); x_50 = l_Lean_Elab_Command_elabMacroRulesAux___closed__6; -lean_inc(x_17); +lean_inc(x_16); x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_17); +lean_ctor_set(x_51, 0, x_16); lean_ctor_set(x_51, 1, x_50); x_52 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; x_53 = l_Lean_Elab_Command_elabMacroRulesAux___closed__8; x_54 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; -lean_inc(x_22); -lean_inc(x_19); -x_55 = l_Lean_addMacroScope(x_19, x_54, x_22); -lean_inc_ref(x_23); -x_56 = l_Lean_Name_mkStr2(x_23, x_52); +lean_inc(x_18); +lean_inc(x_21); +x_55 = l_Lean_addMacroScope(x_21, x_54, x_18); +lean_inc_ref(x_20); +x_56 = l_Lean_Name_mkStr2(x_20, x_52); lean_inc(x_56); x_57 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_57, 0, x_56); @@ -2623,87 +2623,87 @@ lean_ctor_set(x_59, 1, x_45); x_60 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_60, 0, x_57); lean_ctor_set(x_60, 1, x_59); -lean_inc(x_17); +lean_inc(x_16); x_61 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_61, 0, x_17); +lean_ctor_set(x_61, 0, x_16); lean_ctor_set(x_61, 1, x_53); lean_ctor_set(x_61, 2, x_55); lean_ctor_set(x_61, 3, x_60); x_62 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; -lean_inc(x_17); +lean_inc(x_16); x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_17); +lean_ctor_set(x_63, 0, x_16); lean_ctor_set(x_63, 1, x_62); x_64 = l_Lean_Elab_Command_elabMacroRulesAux___closed__11; -lean_inc_ref(x_23); -x_65 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_64); -lean_inc(x_17); +lean_inc_ref(x_20); +x_65 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_64); +lean_inc(x_16); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_17); +lean_ctor_set(x_66, 0, x_16); lean_ctor_set(x_66, 1, x_64); x_67 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; -lean_inc_ref(x_23); -x_68 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_67); +lean_inc_ref(x_20); +x_68 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_67); x_69 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__7; -lean_inc_ref(x_23); -x_70 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_69); -x_71 = l_Array_append___redArg(x_15, x_13); +lean_inc_ref(x_20); +x_70 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_69); +x_71 = l_Array_append___redArg(x_25, x_13); lean_dec(x_13); x_72 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__9; -lean_inc(x_17); +lean_inc(x_16); x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_17); +lean_ctor_set(x_73, 0, x_16); lean_ctor_set(x_73, 1, x_72); x_74 = l_Lean_Elab_Command_elabMacroRulesAux___closed__13; -lean_inc_ref(x_23); -x_75 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_74); +lean_inc_ref(x_20); +x_75 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_74); x_76 = l_Lean_Elab_Command_elabMacroRulesAux___closed__14; -lean_inc(x_17); +lean_inc(x_16); x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_17); +lean_ctor_set(x_77, 0, x_16); lean_ctor_set(x_77, 1, x_76); -lean_inc(x_17); -x_78 = l_Lean_Syntax_node1(x_17, x_75, x_77); -lean_inc(x_25); -lean_inc(x_17); -x_79 = l_Lean_Syntax_node1(x_17, x_25, x_78); -lean_inc(x_25); -lean_inc(x_17); -x_80 = l_Lean_Syntax_node1(x_17, x_25, x_79); +lean_inc(x_16); +x_78 = l_Lean_Syntax_node1(x_16, x_75, x_77); +lean_inc(x_23); +lean_inc(x_16); +x_79 = l_Lean_Syntax_node1(x_16, x_23, x_78); +lean_inc(x_23); +lean_inc(x_16); +x_80 = l_Lean_Syntax_node1(x_16, x_23, x_79); x_81 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__13; -lean_inc(x_17); +lean_inc(x_16); x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_17); +lean_ctor_set(x_82, 0, x_16); lean_ctor_set(x_82, 1, x_81); x_83 = l_Lean_Elab_Command_elabMacroRulesAux___closed__15; -lean_inc_ref(x_23); -x_84 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_83); +lean_inc_ref(x_20); +x_84 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_83); x_85 = l_Lean_Elab_Command_elabMacroRulesAux___closed__16; -lean_inc(x_17); +lean_inc(x_16); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_17); +lean_ctor_set(x_86, 0, x_16); lean_ctor_set(x_86, 1, x_85); x_87 = l_Lean_Elab_Command_elabMacroRulesAux___closed__17; -lean_inc_ref(x_23); -x_88 = l_Lean_Name_mkStr4(x_23, x_29, x_30, x_87); +lean_inc_ref(x_20); +x_88 = l_Lean_Name_mkStr4(x_20, x_29, x_30, x_87); x_89 = l_Lean_Elab_Command_elabMacroRulesAux___closed__19; x_90 = l_Lean_Elab_Command_elabMacroRulesAux___closed__20; -lean_inc(x_22); -lean_inc(x_19); -x_91 = l_Lean_addMacroScope(x_19, x_90, x_22); +lean_inc(x_18); +lean_inc(x_21); +x_91 = l_Lean_addMacroScope(x_21, x_90, x_18); x_92 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; -lean_inc(x_17); +lean_inc(x_16); x_93 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_93, 0, x_17); +lean_ctor_set(x_93, 0, x_16); lean_ctor_set(x_93, 1, x_89); lean_ctor_set(x_93, 2, x_91); lean_ctor_set(x_93, 3, x_92); x_94 = l_Lean_Elab_Command_elabMacroRulesAux___closed__26; x_95 = l_Lean_Elab_Command_elabMacroRulesAux___closed__27; x_96 = l_Lean_Elab_Command_elabMacroRulesAux___closed__28; -x_97 = l_Lean_Name_mkStr4(x_23, x_52, x_95, x_96); +x_97 = l_Lean_Name_mkStr4(x_20, x_52, x_95, x_96); lean_inc(x_97); -x_98 = l_Lean_addMacroScope(x_19, x_97, x_22); +x_98 = l_Lean_addMacroScope(x_21, x_97, x_18); lean_inc(x_97); x_99 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_99, 0, x_97); @@ -2716,35 +2716,35 @@ lean_ctor_set(x_101, 1, x_45); x_102 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_102, 0, x_99); lean_ctor_set(x_102, 1, x_101); -lean_inc(x_17); +lean_inc(x_16); x_103 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_103, 0, x_17); +lean_ctor_set(x_103, 0, x_16); lean_ctor_set(x_103, 1, x_94); lean_ctor_set(x_103, 2, x_98); lean_ctor_set(x_103, 3, x_102); -lean_inc(x_25); -lean_inc(x_17); -x_104 = l_Lean_Syntax_node1(x_17, x_25, x_103); -lean_inc(x_17); -x_105 = l_Lean_Syntax_node2(x_17, x_88, x_93, x_104); -lean_inc(x_17); -x_106 = l_Lean_Syntax_node2(x_17, x_84, x_86, x_105); -lean_inc(x_17); -x_107 = l_Lean_Syntax_node4(x_17, x_70, x_73, x_80, x_82, x_106); +lean_inc(x_23); +lean_inc(x_16); +x_104 = l_Lean_Syntax_node1(x_16, x_23, x_103); +lean_inc(x_16); +x_105 = l_Lean_Syntax_node2(x_16, x_88, x_93, x_104); +lean_inc(x_16); +x_106 = l_Lean_Syntax_node2(x_16, x_84, x_86, x_105); +lean_inc(x_16); +x_107 = l_Lean_Syntax_node4(x_16, x_70, x_73, x_80, x_82, x_106); x_108 = lean_array_push(x_71, x_107); -lean_inc(x_17); +lean_inc(x_16); x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_17); -lean_ctor_set(x_109, 1, x_25); +lean_ctor_set(x_109, 0, x_16); +lean_ctor_set(x_109, 1, x_23); lean_ctor_set(x_109, 2, x_108); -lean_inc(x_17); -x_110 = l_Lean_Syntax_node1(x_17, x_68, x_109); -lean_inc(x_17); -x_111 = l_Lean_Syntax_node2(x_17, x_65, x_66, x_110); +lean_inc(x_16); +x_110 = l_Lean_Syntax_node1(x_16, x_68, x_109); +lean_inc(x_16); +x_111 = l_Lean_Syntax_node2(x_16, x_65, x_66, x_110); x_112 = l_Lean_Elab_Command_elabMacroRulesAux___closed__29; x_113 = lean_array_push(x_112, x_28); x_114 = lean_array_push(x_113, x_40); -x_115 = lean_array_push(x_114, x_24); +x_115 = lean_array_push(x_114, x_17); x_116 = lean_array_push(x_115, x_41); x_117 = lean_array_push(x_116, x_49); x_118 = lean_array_push(x_117, x_51); @@ -2752,8 +2752,8 @@ x_119 = lean_array_push(x_118, x_61); x_120 = lean_array_push(x_119, x_63); x_121 = lean_array_push(x_120, x_111); x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_17); -lean_ctor_set(x_122, 1, x_21); +lean_ctor_set(x_122, 0, x_16); +lean_ctor_set(x_122, 1, x_15); lean_ctor_set(x_122, 2, x_121); if (lean_is_scalar(x_14)) { x_123 = lean_alloc_ctor(0, 1, 0); @@ -2775,17 +2775,17 @@ if (lean_obj_tag(x_1) == 0) { lean_object* x_136; x_136 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; -x_15 = x_135; -x_16 = x_132; -x_17 = x_125; -x_18 = lean_box(0); -x_19 = x_129; -x_20 = x_126; -x_21 = x_133; -x_22 = x_127; -x_23 = x_131; -x_24 = x_128; -x_25 = x_134; +x_15 = x_133; +x_16 = x_125; +x_17 = x_127; +x_18 = x_126; +x_19 = x_132; +x_20 = x_131; +x_21 = x_129; +x_22 = x_128; +x_23 = x_134; +x_24 = lean_box(0); +x_25 = x_135; x_26 = x_136; goto block_124; } @@ -2796,17 +2796,17 @@ x_137 = lean_ctor_get(x_1, 0); lean_inc(x_137); lean_dec_ref(x_1); x_138 = l_Array_mkArray1___redArg(x_137); -x_15 = x_135; -x_16 = x_132; -x_17 = x_125; -x_18 = lean_box(0); -x_19 = x_129; -x_20 = x_126; -x_21 = x_133; -x_22 = x_127; -x_23 = x_131; -x_24 = x_128; -x_25 = x_134; +x_15 = x_133; +x_16 = x_125; +x_17 = x_127; +x_18 = x_126; +x_19 = x_132; +x_20 = x_131; +x_21 = x_129; +x_22 = x_128; +x_23 = x_134; +x_24 = lean_box(0); +x_25 = x_135; x_26 = x_138; goto block_124; } @@ -3523,24 +3523,24 @@ goto block_83; block_36: { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_inc_ref(x_24); -x_29 = l_Array_append___redArg(x_24, x_28); +lean_inc_ref(x_25); +x_29 = l_Array_append___redArg(x_25, x_28); lean_dec_ref(x_28); -lean_inc(x_22); +lean_inc(x_24); lean_inc(x_21); x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_21); -lean_ctor_set(x_30, 1, x_22); +lean_ctor_set(x_30, 1, x_24); lean_ctor_set(x_30, 2, x_29); -x_31 = l_Array_append___redArg(x_24, x_11); +x_31 = l_Array_append___redArg(x_25, x_11); lean_inc(x_21); x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_21); -lean_ctor_set(x_32, 1, x_22); +lean_ctor_set(x_32, 1, x_24); lean_ctor_set(x_32, 2, x_31); lean_inc(x_21); x_33 = l_Lean_Syntax_node1(x_21, x_1, x_32); -x_34 = l_Lean_Syntax_node6(x_21, x_2, x_23, x_25, x_3, x_27, x_30, x_33); +x_34 = l_Lean_Syntax_node6(x_21, x_2, x_23, x_22, x_3, x_27, x_30, x_33); if (lean_is_scalar(x_18)) { x_35 = lean_alloc_ctor(0, 1, 0); } else { @@ -3555,11 +3555,11 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_inc_ref(x_39); x_42 = l_Array_append___redArg(x_39, x_41); lean_dec_ref(x_41); -lean_inc(x_37); +lean_inc(x_38); lean_inc(x_21); x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_21); -lean_ctor_set(x_43, 1, x_37); +lean_ctor_set(x_43, 1, x_38); lean_ctor_set(x_43, 2, x_42); lean_inc(x_21); x_44 = lean_alloc_ctor(2, 2, 0); @@ -3569,10 +3569,10 @@ if (lean_obj_tag(x_10) == 0) { lean_object* x_45; x_45 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; -x_22 = x_37; -x_23 = x_38; -x_24 = x_39; -x_25 = x_43; +x_22 = x_43; +x_23 = x_37; +x_24 = x_38; +x_25 = x_39; x_26 = lean_box(0); x_27 = x_44; x_28 = x_45; @@ -3606,10 +3606,10 @@ x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_21); lean_ctor_set(x_55, 1, x_54); x_56 = l_Array_mkArray5___redArg(x_49, x_51, x_53, x_47, x_55); -x_22 = x_37; -x_23 = x_38; -x_24 = x_39; -x_25 = x_43; +x_22 = x_43; +x_23 = x_37; +x_24 = x_38; +x_25 = x_39; x_26 = lean_box(0); x_27 = x_44; x_28 = x_56; @@ -3635,8 +3635,8 @@ lean_dec_ref(x_8); lean_dec_ref(x_7); lean_dec_ref(x_6); x_64 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; -x_37 = x_58; -x_38 = x_63; +x_37 = x_63; +x_38 = x_58; x_39 = x_59; x_40 = lean_box(0); x_41 = x_64; @@ -3669,8 +3669,8 @@ lean_ctor_set(x_73, 1, x_72); lean_inc(x_21); x_74 = l_Lean_Syntax_node3(x_21, x_67, x_69, x_71, x_73); x_75 = l_Array_mkArray1___redArg(x_74); -x_37 = x_58; -x_38 = x_63; +x_37 = x_63; +x_38 = x_58; x_39 = x_59; x_40 = lean_box(0); x_41 = x_75; @@ -3964,7 +3964,7 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lam__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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; 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_117; lean_object* x_118; 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; 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_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_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_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; x_15 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__4; x_16 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__5; x_17 = l_Lean_Elab_Command_elabMacroRules___lam__1___closed__0; @@ -4058,9 +4058,9 @@ if (x_177 == 0) lean_object* x_178; lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); lean_dec(x_1); x_178 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacroRulesAux_spec__0___redArg(); return x_178; @@ -4087,9 +4087,9 @@ lean_dec(x_182); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); lean_dec(x_1); x_186 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacroRulesAux_spec__0___redArg(); return x_186; @@ -4110,9 +4110,9 @@ lean_dec(x_182); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); x_190 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacroRulesAux_spec__0___redArg(); return x_190; } @@ -4151,7 +4151,7 @@ lean_inc(x_174); x_203 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__0___boxed), 10, 7); lean_closure_set(x_203, 0, x_201); lean_closure_set(x_203, 1, x_202); -lean_closure_set(x_203, 2, x_169); +lean_closure_set(x_203, 2, x_168); lean_closure_set(x_203, 3, x_172); lean_closure_set(x_203, 4, x_174); lean_closure_set(x_203, 5, x_180); @@ -4160,7 +4160,7 @@ if (x_177 == 0) { lean_object* x_204; lean_dec(x_174); -x_204 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_168, x_167); +x_204 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_169, x_170); return x_204; } else @@ -4174,7 +4174,7 @@ if (x_206 == 0) { lean_object* x_207; lean_dec(x_205); -x_207 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_168, x_167); +x_207 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_169, x_170); return x_207; } else @@ -4187,13 +4187,13 @@ x_210 = l_Lean_Syntax_isOfKind(x_208, x_209); if (x_210 == 0) { lean_object* x_211; -x_211 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_168, x_167); +x_211 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_177, x_169, x_170); return x_211; } else { lean_object* x_212; -x_212 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_193, x_168, x_167); +x_212 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_203, x_193, x_169, x_170); return x_212; } } @@ -4233,7 +4233,7 @@ lean_inc(x_174); x_225 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__0___boxed), 10, 7); lean_closure_set(x_225, 0, x_223); lean_closure_set(x_225, 1, x_224); -lean_closure_set(x_225, 2, x_169); +lean_closure_set(x_225, 2, x_168); lean_closure_set(x_225, 3, x_172); lean_closure_set(x_225, 4, x_174); lean_closure_set(x_225, 5, x_180); @@ -4242,7 +4242,7 @@ if (x_177 == 0) { lean_object* x_226; lean_dec(x_174); -x_226 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_168, x_167); +x_226 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_169, x_170); return x_226; } else @@ -4256,7 +4256,7 @@ if (x_228 == 0) { lean_object* x_229; lean_dec(x_227); -x_229 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_168, x_167); +x_229 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_169, x_170); return x_229; } else @@ -4269,13 +4269,13 @@ x_232 = l_Lean_Syntax_isOfKind(x_230, x_231); if (x_232 == 0) { lean_object* x_233; -x_233 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_168, x_167); +x_233 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_177, x_169, x_170); return x_233; } else { lean_object* x_234; -x_234 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_215, x_168, x_167); +x_234 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_225, x_215, x_169, x_170); return x_234; } } @@ -4315,7 +4315,7 @@ lean_inc(x_174); x_246 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__0___boxed), 10, 7); lean_closure_set(x_246, 0, x_244); lean_closure_set(x_246, 1, x_245); -lean_closure_set(x_246, 2, x_169); +lean_closure_set(x_246, 2, x_168); lean_closure_set(x_246, 3, x_172); lean_closure_set(x_246, 4, x_174); lean_closure_set(x_246, 5, x_180); @@ -4324,7 +4324,7 @@ if (x_177 == 0) { lean_object* x_247; lean_dec(x_174); -x_247 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_168, x_167); +x_247 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_169, x_170); return x_247; } else @@ -4338,7 +4338,7 @@ if (x_249 == 0) { lean_object* x_250; lean_dec(x_248); -x_250 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_168, x_167); +x_250 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_169, x_170); return x_250; } else @@ -4351,13 +4351,13 @@ x_253 = l_Lean_Syntax_isOfKind(x_251, x_252); if (x_253 == 0) { lean_object* x_254; -x_254 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_168, x_167); +x_254 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_177, x_169, x_170); return x_254; } else { lean_object* x_255; -x_255 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_236, x_168, x_167); +x_255 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_246, x_236, x_169, x_170); return x_255; } } @@ -4398,7 +4398,7 @@ lean_inc(x_174); x_267 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__0___boxed), 10, 7); lean_closure_set(x_267, 0, x_265); lean_closure_set(x_267, 1, x_266); -lean_closure_set(x_267, 2, x_169); +lean_closure_set(x_267, 2, x_168); lean_closure_set(x_267, 3, x_172); lean_closure_set(x_267, 4, x_174); lean_closure_set(x_267, 5, x_180); @@ -4407,7 +4407,7 @@ if (x_177 == 0) { lean_object* x_268; lean_dec(x_174); -x_268 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_168, x_167); +x_268 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_169, x_170); return x_268; } else @@ -4421,7 +4421,7 @@ if (x_270 == 0) { lean_object* x_271; lean_dec(x_269); -x_271 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_168, x_167); +x_271 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_169, x_170); return x_271; } else @@ -4434,13 +4434,13 @@ x_274 = l_Lean_Syntax_isOfKind(x_272, x_273); if (x_274 == 0) { lean_object* x_275; -x_275 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_168, x_167); +x_275 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_177, x_169, x_170); return x_275; } else { lean_object* x_276; -x_276 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_257, x_168, x_167); +x_276 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_267, x_257, x_169, x_170); return x_276; } } @@ -4482,7 +4482,7 @@ lean_inc(x_174); x_289 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__0___boxed), 10, 7); lean_closure_set(x_289, 0, x_287); lean_closure_set(x_289, 1, x_288); -lean_closure_set(x_289, 2, x_169); +lean_closure_set(x_289, 2, x_168); lean_closure_set(x_289, 3, x_172); lean_closure_set(x_289, 4, x_174); lean_closure_set(x_289, 5, x_180); @@ -4490,12 +4490,12 @@ lean_closure_set(x_289, 6, x_280); if (x_177 == 0) { lean_dec(x_174); -x_5 = x_279; -x_6 = x_257; -x_7 = lean_box(0); -x_8 = x_289; -x_9 = x_168; -x_10 = x_167; +x_5 = lean_box(0); +x_6 = x_289; +x_7 = x_169; +x_8 = x_257; +x_9 = x_279; +x_10 = x_170; x_11 = x_279; goto block_14; } @@ -4509,12 +4509,12 @@ x_291 = l_Lean_Syntax_matchesNull(x_290, x_171); if (x_291 == 0) { lean_dec(x_290); -x_5 = x_279; -x_6 = x_257; -x_7 = lean_box(0); -x_8 = x_289; -x_9 = x_168; -x_10 = x_167; +x_5 = lean_box(0); +x_6 = x_289; +x_7 = x_169; +x_8 = x_257; +x_9 = x_279; +x_10 = x_170; x_11 = x_279; goto block_14; } @@ -4527,19 +4527,19 @@ x_293 = l_Lean_Elab_Command_elabMacroRules___lam__1___closed__13; x_294 = l_Lean_Syntax_isOfKind(x_292, x_293); if (x_294 == 0) { -x_5 = x_279; -x_6 = x_257; -x_7 = lean_box(0); -x_8 = x_289; -x_9 = x_168; -x_10 = x_167; +x_5 = lean_box(0); +x_6 = x_289; +x_7 = x_169; +x_8 = x_257; +x_9 = x_279; +x_10 = x_170; x_11 = x_279; goto block_14; } else { lean_object* x_295; -x_295 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_289, x_279, x_168, x_167); +x_295 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_289, x_279, x_169, x_170); return x_295; } } @@ -4549,19 +4549,19 @@ else { lean_object* x_296; lean_dec(x_192); -x_296 = l_Lean_Elab_Command_getRef___redArg(x_168); +x_296 = l_Lean_Elab_Command_getRef___redArg(x_169); if (lean_obj_tag(x_296) == 0) { lean_object* x_297; uint8_t x_298; x_297 = lean_ctor_get(x_296, 0); lean_inc(x_297); lean_dec_ref(x_296); -x_298 = !lean_is_exclusive(x_168); +x_298 = !lean_is_exclusive(x_169); if (x_298 == 0) { lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; -x_299 = lean_ctor_get(x_168, 5); -x_300 = lean_ctor_get(x_168, 7); +x_299 = lean_ctor_get(x_169, 5); +x_300 = lean_ctor_get(x_169, 7); lean_dec(x_300); x_301 = l_Lean_Syntax_getArg(x_213, x_179); lean_dec(x_213); @@ -4580,15 +4580,15 @@ x_308 = l_Lean_replaceRef(x_307, x_297); lean_dec(x_297); lean_dec_ref(x_307); lean_inc(x_299); -lean_ctor_set(x_168, 7, x_308); -x_309 = l_Lean_Elab_Command_getRef___redArg(x_168); +lean_ctor_set(x_169, 7, x_308); +x_309 = l_Lean_Elab_Command_getRef___redArg(x_169); if (lean_obj_tag(x_309) == 0) { lean_object* x_310; lean_object* x_311; x_310 = lean_ctor_get(x_309, 0); lean_inc(x_310); lean_dec_ref(x_309); -x_311 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_168); +x_311 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_169); if (lean_obj_tag(x_311) == 0) { lean_object* x_312; @@ -4598,43 +4598,43 @@ lean_dec(x_310); if (lean_obj_tag(x_299) == 0) { lean_object* x_313; -x_313 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_167); +x_313 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_170); lean_dec_ref(x_313); -x_132 = x_183; -x_133 = x_305; -x_134 = x_169; -x_135 = x_301; -x_136 = x_191; -x_137 = x_175; -x_138 = x_168; -x_139 = x_299; -x_140 = x_174; -x_141 = x_312; -x_142 = x_277; -x_143 = x_171; -x_144 = x_180; -x_145 = x_172; -x_146 = x_167; +x_132 = x_168; +x_133 = x_277; +x_134 = x_172; +x_135 = x_183; +x_136 = x_305; +x_137 = x_180; +x_138 = x_169; +x_139 = x_175; +x_140 = x_299; +x_141 = x_301; +x_142 = x_171; +x_143 = x_170; +x_144 = x_312; +x_145 = x_174; +x_146 = x_191; x_147 = lean_box(0); goto block_164; } else { -x_132 = x_183; -x_133 = x_305; -x_134 = x_169; -x_135 = x_301; -x_136 = x_191; -x_137 = x_175; -x_138 = x_168; -x_139 = x_299; -x_140 = x_174; -x_141 = x_312; -x_142 = x_277; -x_143 = x_171; -x_144 = x_180; -x_145 = x_172; -x_146 = x_167; +x_132 = x_168; +x_133 = x_277; +x_134 = x_172; +x_135 = x_183; +x_136 = x_305; +x_137 = x_180; +x_138 = x_169; +x_139 = x_175; +x_140 = x_299; +x_141 = x_301; +x_142 = x_171; +x_143 = x_170; +x_144 = x_312; +x_145 = x_174; +x_146 = x_191; x_147 = lean_box(0); goto block_164; } @@ -4643,7 +4643,7 @@ else { uint8_t x_314; lean_dec(x_310); -lean_dec_ref(x_168); +lean_dec_ref(x_169); lean_dec(x_301); lean_dec(x_299); lean_dec(x_277); @@ -4651,8 +4651,8 @@ lean_dec(x_191); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec(x_167); +lean_dec(x_170); +lean_dec(x_168); x_314 = !lean_is_exclusive(x_311); if (x_314 == 0) { @@ -4672,7 +4672,7 @@ return x_316; } else { -lean_dec_ref(x_168); +lean_dec_ref(x_169); lean_dec(x_301); lean_dec(x_299); lean_dec(x_277); @@ -4680,24 +4680,24 @@ lean_dec(x_191); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec(x_167); +lean_dec(x_170); +lean_dec(x_168); return x_309; } } else { lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; uint8_t x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; -x_317 = lean_ctor_get(x_168, 0); -x_318 = lean_ctor_get(x_168, 1); -x_319 = lean_ctor_get(x_168, 2); -x_320 = lean_ctor_get(x_168, 3); -x_321 = lean_ctor_get(x_168, 4); -x_322 = lean_ctor_get(x_168, 5); -x_323 = lean_ctor_get(x_168, 6); -x_324 = lean_ctor_get(x_168, 8); -x_325 = lean_ctor_get(x_168, 9); -x_326 = lean_ctor_get_uint8(x_168, sizeof(void*)*10); +x_317 = lean_ctor_get(x_169, 0); +x_318 = lean_ctor_get(x_169, 1); +x_319 = lean_ctor_get(x_169, 2); +x_320 = lean_ctor_get(x_169, 3); +x_321 = lean_ctor_get(x_169, 4); +x_322 = lean_ctor_get(x_169, 5); +x_323 = lean_ctor_get(x_169, 6); +x_324 = lean_ctor_get(x_169, 8); +x_325 = lean_ctor_get(x_169, 9); +x_326 = lean_ctor_get_uint8(x_169, sizeof(void*)*10); lean_inc(x_325); lean_inc(x_324); lean_inc(x_323); @@ -4707,7 +4707,7 @@ lean_inc(x_320); lean_inc(x_319); lean_inc(x_318); lean_inc(x_317); -lean_dec(x_168); +lean_dec(x_169); x_327 = l_Lean_Syntax_getArg(x_213, x_179); lean_dec(x_213); x_328 = l_Lean_Elab_Command_elabMacroRules___lam__1___closed__11; @@ -4754,43 +4754,43 @@ lean_dec(x_337); if (lean_obj_tag(x_322) == 0) { lean_object* x_340; -x_340 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_167); +x_340 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_170); lean_dec_ref(x_340); -x_132 = x_183; -x_133 = x_331; -x_134 = x_169; -x_135 = x_327; -x_136 = x_191; -x_137 = x_175; +x_132 = x_168; +x_133 = x_277; +x_134 = x_172; +x_135 = x_183; +x_136 = x_331; +x_137 = x_180; x_138 = x_335; -x_139 = x_322; -x_140 = x_174; -x_141 = x_339; -x_142 = x_277; -x_143 = x_171; -x_144 = x_180; -x_145 = x_172; -x_146 = x_167; +x_139 = x_175; +x_140 = x_322; +x_141 = x_327; +x_142 = x_171; +x_143 = x_170; +x_144 = x_339; +x_145 = x_174; +x_146 = x_191; x_147 = lean_box(0); goto block_164; } else { -x_132 = x_183; -x_133 = x_331; -x_134 = x_169; -x_135 = x_327; -x_136 = x_191; -x_137 = x_175; +x_132 = x_168; +x_133 = x_277; +x_134 = x_172; +x_135 = x_183; +x_136 = x_331; +x_137 = x_180; x_138 = x_335; -x_139 = x_322; -x_140 = x_174; -x_141 = x_339; -x_142 = x_277; -x_143 = x_171; -x_144 = x_180; -x_145 = x_172; -x_146 = x_167; +x_139 = x_175; +x_140 = x_322; +x_141 = x_327; +x_142 = x_171; +x_143 = x_170; +x_144 = x_339; +x_145 = x_174; +x_146 = x_191; x_147 = lean_box(0); goto block_164; } @@ -4807,8 +4807,8 @@ lean_dec(x_191); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec(x_167); +lean_dec(x_170); +lean_dec(x_168); x_341 = lean_ctor_get(x_338, 0); lean_inc(x_341); if (lean_is_exclusive(x_338)) { @@ -4837,8 +4837,8 @@ lean_dec(x_191); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec(x_167); +lean_dec(x_170); +lean_dec(x_168); return x_336; } } @@ -4851,9 +4851,9 @@ lean_dec(x_191); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); return x_296; } } @@ -4881,27 +4881,27 @@ lean_dec(x_345); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); x_348 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabMacroRulesAux_spec__0___redArg(); return x_348; } else { lean_object* x_349; -x_349 = l_Lean_Elab_Command_getRef___redArg(x_168); +x_349 = l_Lean_Elab_Command_getRef___redArg(x_169); if (lean_obj_tag(x_349) == 0) { lean_object* x_350; uint8_t x_351; x_350 = lean_ctor_get(x_349, 0); lean_inc(x_350); lean_dec_ref(x_349); -x_351 = !lean_is_exclusive(x_168); +x_351 = !lean_is_exclusive(x_169); if (x_351 == 0) { lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; -x_352 = lean_ctor_get(x_168, 7); +x_352 = lean_ctor_get(x_169, 7); lean_dec(x_352); x_353 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__5___boxed), 14, 9); lean_closure_set(x_353, 0, x_346); @@ -4912,7 +4912,7 @@ lean_closure_set(x_353, 4, x_172); lean_closure_set(x_353, 5, x_15); lean_closure_set(x_353, 6, x_16); lean_closure_set(x_353, 7, x_175); -lean_closure_set(x_353, 8, x_169); +lean_closure_set(x_353, 8, x_168); x_354 = l_Lean_Syntax_getArg(x_345, x_166); lean_dec(x_345); x_355 = l_Lean_Syntax_getArgs(x_354); @@ -4934,8 +4934,8 @@ lean_ctor_set(x_362, 2, x_361); x_363 = l_Lean_replaceRef(x_362, x_350); lean_dec(x_350); lean_dec_ref(x_362); -lean_ctor_set(x_168, 7, x_363); -x_364 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_355, x_17, x_353, x_168, x_167); +lean_ctor_set(x_169, 7, x_363); +x_364 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_355, x_17, x_353, x_169, x_170); lean_dec_ref(x_355); if (lean_obj_tag(x_364) == 0) { @@ -4979,16 +4979,16 @@ return x_370; else { lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; uint8_t x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; -x_371 = lean_ctor_get(x_168, 0); -x_372 = lean_ctor_get(x_168, 1); -x_373 = lean_ctor_get(x_168, 2); -x_374 = lean_ctor_get(x_168, 3); -x_375 = lean_ctor_get(x_168, 4); -x_376 = lean_ctor_get(x_168, 5); -x_377 = lean_ctor_get(x_168, 6); -x_378 = lean_ctor_get(x_168, 8); -x_379 = lean_ctor_get(x_168, 9); -x_380 = lean_ctor_get_uint8(x_168, sizeof(void*)*10); +x_371 = lean_ctor_get(x_169, 0); +x_372 = lean_ctor_get(x_169, 1); +x_373 = lean_ctor_get(x_169, 2); +x_374 = lean_ctor_get(x_169, 3); +x_375 = lean_ctor_get(x_169, 4); +x_376 = lean_ctor_get(x_169, 5); +x_377 = lean_ctor_get(x_169, 6); +x_378 = lean_ctor_get(x_169, 8); +x_379 = lean_ctor_get(x_169, 9); +x_380 = lean_ctor_get_uint8(x_169, sizeof(void*)*10); lean_inc(x_379); lean_inc(x_378); lean_inc(x_377); @@ -4998,7 +4998,7 @@ lean_inc(x_374); lean_inc(x_373); lean_inc(x_372); lean_inc(x_371); -lean_dec(x_168); +lean_dec(x_169); x_381 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lam__5___boxed), 14, 9); lean_closure_set(x_381, 0, x_346); lean_closure_set(x_381, 1, x_18); @@ -5008,7 +5008,7 @@ lean_closure_set(x_381, 4, x_172); lean_closure_set(x_381, 5, x_15); lean_closure_set(x_381, 6, x_16); lean_closure_set(x_381, 7, x_175); -lean_closure_set(x_381, 8, x_169); +lean_closure_set(x_381, 8, x_168); x_382 = l_Lean_Syntax_getArg(x_345, x_166); lean_dec(x_345); x_383 = l_Lean_Syntax_getArgs(x_382); @@ -5042,7 +5042,7 @@ lean_ctor_set(x_392, 7, x_391); lean_ctor_set(x_392, 8, x_378); lean_ctor_set(x_392, 9, x_379); lean_ctor_set_uint8(x_392, sizeof(void*)*10, x_380); -x_393 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_383, x_17, x_381, x_392, x_167); +x_393 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_383, x_17, x_381, x_392, x_170); lean_dec_ref(x_383); if (lean_obj_tag(x_393) == 0) { @@ -5092,9 +5092,9 @@ lean_dec(x_345); lean_dec(x_180); lean_dec(x_174); lean_dec(x_172); -lean_dec(x_169); -lean_dec_ref(x_168); -lean_dec(x_167); +lean_dec(x_170); +lean_dec_ref(x_169); +lean_dec(x_168); return x_349; } } @@ -5151,10 +5151,10 @@ x_415 = l_Lean_Syntax_getArgs(x_414); lean_dec(x_414); x_416 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_416, 0, x_415); -x_167 = x_403; -x_168 = x_402; -x_169 = x_401; -x_170 = lean_box(0); +x_167 = lean_box(0); +x_168 = x_401; +x_169 = x_402; +x_170 = x_403; x_171 = x_405; x_172 = x_416; goto block_400; @@ -5166,10 +5166,10 @@ else lean_object* x_417; lean_dec(x_406); x_417 = lean_box(0); -x_167 = x_403; -x_168 = x_402; -x_169 = x_401; -x_170 = lean_box(0); +x_167 = lean_box(0); +x_168 = x_401; +x_169 = x_402; +x_170 = x_403; x_171 = x_405; x_172 = x_417; goto block_400; @@ -5181,115 +5181,115 @@ goto block_400; if (x_11 == 0) { lean_object* x_12; -x_12 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_8, x_6, x_9, x_10); +x_12 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_6, x_8, x_7, x_10); return x_12; } else { lean_object* x_13; -x_13 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_8, x_5, x_9, x_10); +x_13 = l_Lean_withExporting___at___00Lean_Elab_Command_elabMacroRules_spec__0___redArg(x_6, x_9, x_7, x_10); return x_13; } } block_84: { lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_inc_ref(x_26); -x_36 = l_Array_append___redArg(x_26, x_35); +lean_inc_ref(x_30); +x_36 = l_Array_append___redArg(x_30, x_35); lean_dec_ref(x_35); lean_inc(x_21); -lean_inc(x_28); +lean_inc(x_25); x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_28); +lean_ctor_set(x_37, 0, x_25); lean_ctor_set(x_37, 1, x_21); lean_ctor_set(x_37, 2, x_36); x_38 = l_Lean_Elab_Command_elabMacroRulesAux___closed__0; -lean_inc_ref(x_23); -x_39 = l_Lean_Name_mkStr4(x_15, x_16, x_23, x_38); +lean_inc_ref(x_28); +x_39 = l_Lean_Name_mkStr4(x_15, x_16, x_28, x_38); x_40 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; -lean_inc(x_28); +lean_inc(x_25); x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_28); +lean_ctor_set(x_41, 0, x_25); lean_ctor_set(x_41, 1, x_40); -lean_inc_ref(x_26); -x_42 = l_Array_append___redArg(x_26, x_20); -lean_dec_ref(x_20); +lean_inc_ref(x_30); +x_42 = l_Array_append___redArg(x_30, x_33); +lean_dec_ref(x_33); lean_inc(x_21); -lean_inc(x_28); +lean_inc(x_25); x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_28); +lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_21); lean_ctor_set(x_43, 2, x_42); x_44 = l_Lean_Elab_Command_elabMacroRulesAux___closed__2; -lean_inc(x_28); +lean_inc(x_25); x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_28); +lean_ctor_set(x_45, 0, x_25); lean_ctor_set(x_45, 1, x_44); -lean_inc(x_28); -x_46 = l_Lean_Syntax_node3(x_28, x_39, x_41, x_43, x_45); +lean_inc(x_25); +x_46 = l_Lean_Syntax_node3(x_25, x_39, x_41, x_43, x_45); lean_inc(x_21); -lean_inc(x_28); -x_47 = l_Lean_Syntax_node1(x_28, x_21, x_46); -lean_inc(x_28); +lean_inc(x_25); +x_47 = l_Lean_Syntax_node1(x_25, x_21, x_46); +lean_inc(x_25); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_28); -lean_ctor_set(x_48, 1, x_22); -x_49 = l_Lean_TSyntax_getId(x_24); -x_50 = l_Lean_mkIdentFrom(x_33, x_49, x_19); -lean_dec(x_33); +lean_ctor_set(x_48, 0, x_25); +lean_ctor_set(x_48, 1, x_29); +x_49 = l_Lean_TSyntax_getId(x_34); +x_50 = l_Lean_mkIdentFrom(x_22, x_49, x_19); +lean_dec(x_22); lean_inc(x_21); -lean_inc(x_28); -x_51 = l_Lean_Syntax_node2(x_28, x_21, x_50, x_24); +lean_inc(x_25); +x_51 = l_Lean_Syntax_node2(x_25, x_21, x_50, x_34); x_52 = l_Lean_Elab_Command_elabMacroRulesAux___closed__6; -lean_inc(x_28); +lean_inc(x_25); x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_28); +lean_ctor_set(x_53, 0, x_25); lean_ctor_set(x_53, 1, x_52); x_54 = l_Lean_Elab_Command_elabMacroRulesAux___closed__8; x_55 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; -x_56 = l_Lean_addMacroScope(x_30, x_55, x_34); +x_56 = l_Lean_addMacroScope(x_26, x_55, x_23); x_57 = l_Lean_Elab_Command_elabMacroRules___lam__1___closed__6; -lean_inc(x_28); +lean_inc(x_25); x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_28); +lean_ctor_set(x_58, 0, x_25); lean_ctor_set(x_58, 1, x_54); lean_ctor_set(x_58, 2, x_56); lean_ctor_set(x_58, 3, x_57); x_59 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; -lean_inc(x_28); +lean_inc(x_25); x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_28); +lean_ctor_set(x_60, 0, x_25); lean_ctor_set(x_60, 1, x_59); x_61 = l_Lean_Elab_Command_elabMacroRulesAux___closed__11; -lean_inc_ref(x_23); -x_62 = l_Lean_Name_mkStr4(x_15, x_16, x_23, x_61); -lean_inc(x_28); +lean_inc_ref(x_28); +x_62 = l_Lean_Name_mkStr4(x_15, x_16, x_28, x_61); +lean_inc(x_25); x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_28); +lean_ctor_set(x_63, 0, x_25); lean_ctor_set(x_63, 1, x_61); x_64 = l_Lean_Elab_Command_elabMacroRules___lam__1___closed__7; -x_65 = l_Lean_Name_mkStr4(x_15, x_16, x_23, x_64); +x_65 = l_Lean_Name_mkStr4(x_15, x_16, x_28, x_64); lean_inc(x_21); -lean_inc(x_28); -x_66 = l_Lean_Syntax_node1(x_28, x_21, x_31); -lean_inc(x_28); +lean_inc(x_25); +x_66 = l_Lean_Syntax_node1(x_25, x_21, x_20); +lean_inc(x_25); x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_28); +lean_ctor_set(x_67, 0, x_25); lean_ctor_set(x_67, 1, x_21); -lean_ctor_set(x_67, 2, x_26); +lean_ctor_set(x_67, 2, x_30); x_68 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__13; -lean_inc(x_28); +lean_inc(x_25); x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_28); +lean_ctor_set(x_69, 0, x_25); lean_ctor_set(x_69, 1, x_68); -lean_inc(x_28); -x_70 = l_Lean_Syntax_node4(x_28, x_65, x_66, x_67, x_69, x_25); -lean_inc(x_28); -x_71 = l_Lean_Syntax_node2(x_28, x_62, x_63, x_70); +lean_inc(x_25); +x_70 = l_Lean_Syntax_node4(x_25, x_65, x_66, x_67, x_69, x_31); +lean_inc(x_25); +x_71 = l_Lean_Syntax_node2(x_25, x_62, x_63, x_70); x_72 = l_Lean_Elab_Command_elabMacroRulesAux___closed__29; x_73 = lean_array_push(x_72, x_37); x_74 = lean_array_push(x_73, x_47); -x_75 = lean_array_push(x_74, x_27); +x_75 = lean_array_push(x_74, x_24); x_76 = lean_array_push(x_75, x_48); x_77 = lean_array_push(x_76, x_51); x_78 = lean_array_push(x_77, x_53); @@ -5297,8 +5297,8 @@ x_79 = lean_array_push(x_78, x_58); x_80 = lean_array_push(x_79, x_60); x_81 = lean_array_push(x_80, x_71); x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_28); -lean_ctor_set(x_82, 1, x_32); +lean_ctor_set(x_82, 0, x_25); +lean_ctor_set(x_82, 1, x_27); lean_ctor_set(x_82, 2, x_81); x_83 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_83, 0, x_82); @@ -5310,50 +5310,50 @@ lean_object* x_98; lean_object* x_99; lean_object* x_100; x_98 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; x_99 = l_Lean_Elab_Command_elabMacroRulesAux___closed__33; x_100 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabMacroRulesAux_spec__9_spec__9___closed__12; -if (lean_obj_tag(x_88) == 0) +if (lean_obj_tag(x_85) == 0) { lean_object* x_101; x_101 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; -x_20 = x_85; -x_21 = x_86; -x_22 = x_98; -x_23 = x_91; -x_24 = x_90; -x_25 = x_89; -x_26 = x_100; -x_27 = x_93; -x_28 = x_87; -x_29 = lean_box(0); -x_30 = x_96; -x_31 = x_92; -x_32 = x_99; -x_33 = x_95; -x_34 = x_94; +x_20 = x_86; +x_21 = x_87; +x_22 = x_88; +x_23 = x_89; +x_24 = x_91; +x_25 = x_92; +x_26 = x_96; +x_27 = x_99; +x_28 = x_95; +x_29 = x_98; +x_30 = x_100; +x_31 = x_90; +x_32 = lean_box(0); +x_33 = x_94; +x_34 = x_93; x_35 = x_101; goto block_84; } else { lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_88, 0); +x_102 = lean_ctor_get(x_85, 0); lean_inc(x_102); -lean_dec_ref(x_88); +lean_dec_ref(x_85); x_103 = l_Array_mkArray1___redArg(x_102); -x_20 = x_85; -x_21 = x_86; -x_22 = x_98; -x_23 = x_91; -x_24 = x_90; -x_25 = x_89; -x_26 = x_100; -x_27 = x_93; -x_28 = x_87; -x_29 = lean_box(0); -x_30 = x_96; -x_31 = x_92; -x_32 = x_99; -x_33 = x_95; -x_34 = x_94; +x_20 = x_86; +x_21 = x_87; +x_22 = x_88; +x_23 = x_89; +x_24 = x_91; +x_25 = x_92; +x_26 = x_96; +x_27 = x_99; +x_28 = x_95; +x_29 = x_98; +x_30 = x_100; +x_31 = x_90; +x_32 = lean_box(0); +x_33 = x_94; +x_34 = x_93; x_35 = x_103; goto block_84; } @@ -5376,28 +5376,28 @@ lean_object* x_122; lean_object* x_123; lean_object* x_124; x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); lean_dec_ref(x_121); -x_123 = l_Lean_Parser_Command_visibility_ofAttrKind(x_113); -x_124 = l_Lean_SourceInfo_fromRef(x_120, x_105); +x_123 = l_Lean_Parser_Command_visibility_ofAttrKind(x_116); +x_124 = l_Lean_SourceInfo_fromRef(x_120, x_108); lean_dec(x_120); -if (lean_obj_tag(x_112) == 0) +if (lean_obj_tag(x_113) == 0) { lean_object* x_125; lean_object* x_126; -x_125 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_117); -lean_dec(x_117); +x_125 = l_Lean_getMainModule___at___00Lean_Elab_Command_elabMacroRulesAux_spec__8___redArg(x_115); +lean_dec(x_115); x_126 = lean_ctor_get(x_125, 0); lean_inc(x_126); lean_dec_ref(x_125); -x_85 = x_118; -x_86 = x_106; -x_87 = x_124; -x_88 = x_107; -x_89 = x_108; -x_90 = x_109; -x_91 = x_110; -x_92 = x_115; -x_93 = x_123; -x_94 = x_122; -x_95 = x_116; +x_85 = x_106; +x_86 = x_107; +x_87 = x_109; +x_88 = x_110; +x_89 = x_122; +x_90 = x_114; +x_91 = x_123; +x_92 = x_124; +x_93 = x_117; +x_94 = x_118; +x_95 = x_112; x_96 = x_126; x_97 = lean_box(0); goto block_104; @@ -5405,21 +5405,21 @@ goto block_104; else { lean_object* x_127; -lean_dec(x_117); -x_127 = lean_ctor_get(x_112, 0); +lean_dec(x_115); +x_127 = lean_ctor_get(x_113, 0); lean_inc(x_127); -lean_dec_ref(x_112); -x_85 = x_118; -x_86 = x_106; -x_87 = x_124; -x_88 = x_107; -x_89 = x_108; -x_90 = x_109; -x_91 = x_110; -x_92 = x_115; -x_93 = x_123; -x_94 = x_122; -x_95 = x_116; +lean_dec_ref(x_113); +x_85 = x_106; +x_86 = x_107; +x_87 = x_109; +x_88 = x_110; +x_89 = x_122; +x_90 = x_114; +x_91 = x_123; +x_92 = x_124; +x_93 = x_117; +x_94 = x_118; +x_95 = x_112; x_96 = x_127; x_97 = lean_box(0); goto block_104; @@ -5433,11 +5433,11 @@ lean_dec_ref(x_118); lean_dec(x_117); lean_dec(x_116); lean_dec(x_115); +lean_dec(x_114); lean_dec(x_113); -lean_dec(x_112); -lean_dec_ref(x_110); +lean_dec_ref(x_112); +lean_dec(x_110); lean_dec(x_109); -lean_dec(x_108); lean_dec(x_107); lean_dec(x_106); x_128 = !lean_is_exclusive(x_121); @@ -5463,12 +5463,12 @@ lean_dec_ref(x_118); lean_dec(x_117); lean_dec(x_116); lean_dec(x_115); +lean_dec(x_114); lean_dec(x_113); -lean_dec(x_112); +lean_dec_ref(x_112); lean_dec_ref(x_111); -lean_dec_ref(x_110); +lean_dec(x_110); lean_dec(x_109); -lean_dec(x_108); lean_dec(x_107); lean_dec(x_106); return x_119; @@ -5478,39 +5478,39 @@ return x_119; { lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; x_148 = l_Lean_Elab_Command_elabMacroRulesAux___closed__35; -lean_inc_ref(x_137); -x_149 = l_Lean_Name_mkStr4(x_15, x_16, x_137, x_148); +lean_inc_ref(x_139); +x_149 = l_Lean_Name_mkStr4(x_15, x_16, x_139, x_148); x_150 = l_Lean_Elab_Command_elabMacroRulesAux___closed__38; x_151 = l_Lean_Elab_Command_elabMacroRulesAux___closed__39; -lean_inc(x_141); +lean_inc(x_144); x_152 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_152, 0, x_141); +lean_ctor_set(x_152, 0, x_144); lean_ctor_set(x_152, 1, x_150); -lean_inc(x_136); -lean_inc(x_141); -x_153 = l_Lean_Syntax_node2(x_141, x_151, x_152, x_136); -lean_inc(x_140); -x_154 = l_Lean_Syntax_node2(x_141, x_149, x_140, x_153); -if (lean_obj_tag(x_145) == 0) +lean_inc(x_146); +lean_inc(x_144); +x_153 = l_Lean_Syntax_node2(x_144, x_151, x_152, x_146); +lean_inc(x_145); +x_154 = l_Lean_Syntax_node2(x_144, x_149, x_145, x_153); +if (lean_obj_tag(x_134) == 0) { lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; x_155 = l_Lean_Elab_Command_elabMacroRulesAux___closed__40; -x_156 = lean_mk_empty_array_with_capacity(x_143); +x_156 = lean_mk_empty_array_with_capacity(x_142); x_157 = lean_array_push(x_156, x_154); x_158 = l_Lean_Syntax_SepArray_ofElems(x_155, x_157); lean_dec_ref(x_157); -x_105 = x_132; -x_106 = x_133; -x_107 = x_134; +x_105 = lean_box(0); +x_106 = x_132; +x_107 = x_133; x_108 = x_135; x_109 = x_136; x_110 = x_137; x_111 = x_138; x_112 = x_139; x_113 = x_140; -x_114 = lean_box(0); -x_115 = x_142; -x_116 = x_144; +x_114 = x_141; +x_115 = x_143; +x_116 = x_145; x_117 = x_146; x_118 = x_158; goto block_131; @@ -5518,27 +5518,27 @@ goto block_131; else { lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_159 = lean_ctor_get(x_145, 0); +x_159 = lean_ctor_get(x_134, 0); lean_inc(x_159); -lean_dec_ref(x_145); +lean_dec_ref(x_134); x_160 = l_Lean_Elab_Command_elabMacroRulesAux___closed__40; x_161 = l_Lean_Syntax_TSepArray_getElems___redArg(x_159); lean_dec(x_159); x_162 = lean_array_push(x_161, x_154); x_163 = l_Lean_Syntax_SepArray_ofElems(x_160, x_162); lean_dec_ref(x_162); -x_105 = x_132; -x_106 = x_133; -x_107 = x_134; +x_105 = lean_box(0); +x_106 = x_132; +x_107 = x_133; x_108 = x_135; x_109 = x_136; x_110 = x_137; x_111 = x_138; x_112 = x_139; x_113 = x_140; -x_114 = lean_box(0); -x_115 = x_142; -x_116 = x_144; +x_114 = x_141; +x_115 = x_143; +x_116 = x_145; x_117 = x_146; x_118 = x_163; goto block_131; diff --git a/stage0/stdlib/Lean/Elab/Notation.c b/stage0/stdlib/Lean/Elab/Notation.c index ba13e9849fe1..ed4619f690de 100644 --- a/stage0/stdlib/Lean/Elab/Notation.c +++ b/stage0/stdlib/Lean/Elab/Notation.c @@ -8970,7 +8970,7 @@ return x_86; } else { -lean_object* x_87; size_t x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_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_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; uint8_t 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_196; size_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; uint8_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_233; size_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_object* x_242; lean_object* x_243; uint8_t 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_265; size_t 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; uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; size_t x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; uint8_t 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_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_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_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_465; uint8_t x_466; +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; size_t x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_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; size_t 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; 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; size_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; 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_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; size_t x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; lean_object* x_247; 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_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; size_t 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; 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_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; size_t x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_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_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_465; uint8_t x_466; x_87 = lean_unsigned_to_nat(0u); x_465 = l_Lean_Syntax_getArg(x_1, x_87); x_466 = l_Lean_Syntax_isNone(x_465); @@ -9035,74 +9035,74 @@ goto block_464; block_128: { size_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_103 = lean_array_size(x_98); +x_103 = lean_array_size(x_93); x_104 = lean_box_usize(x_103); -x_105 = lean_box_usize(x_88); +x_105 = lean_box_usize(x_94); x_106 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabNotation_spec__22___boxed), 5, 3); lean_closure_set(x_106, 0, x_104); lean_closure_set(x_106, 1, x_105); -lean_closure_set(x_106, 2, x_98); -lean_inc_ref(x_99); -x_107 = l_Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabNotation_spec__1___redArg(x_106, x_99, x_92); +lean_closure_set(x_106, 2, x_93); +lean_inc_ref(x_98); +x_107 = l_Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabNotation_spec__1___redArg(x_106, x_98, x_92); if (lean_obj_tag(x_107) == 0) { lean_object* x_108; lean_object* x_109; x_108 = lean_ctor_get(x_107, 0); lean_inc(x_108); lean_dec_ref(x_107); -x_109 = l_Lean_Elab_Command_getRef___redArg(x_99); +x_109 = l_Lean_Elab_Command_getRef___redArg(x_98); if (lean_obj_tag(x_109) == 0) { lean_object* x_110; lean_object* x_111; x_110 = lean_ctor_get(x_109, 0); lean_inc(x_110); lean_dec_ref(x_109); -x_111 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_99); +x_111 = l_Lean_Elab_Command_getCurrMacroScope___redArg(x_98); if (lean_obj_tag(x_111) == 0) { lean_object* x_112; size_t x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_dec_ref(x_111); -x_112 = lean_ctor_get(x_99, 5); +x_112 = lean_ctor_get(x_98, 5); x_113 = lean_array_size(x_102); -x_114 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabNotation_spec__23(x_87, x_113, x_88, x_102); -x_115 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote(x_114, x_89); +x_114 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabNotation_spec__23(x_87, x_113, x_94, x_102); +x_115 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote(x_114, x_95); lean_dec_ref(x_114); x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_95); -lean_ctor_set(x_116, 1, x_94); +lean_ctor_set(x_116, 0, x_89); +lean_ctor_set(x_116, 1, x_100); lean_ctor_set(x_116, 2, x_108); -x_117 = l_Lean_SourceInfo_fromRef(x_110, x_97); +x_117 = l_Lean_SourceInfo_fromRef(x_110, x_96); lean_dec(x_110); if (lean_obj_tag(x_112) == 0) { lean_object* x_118; x_118 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabNotation_spec__1_spec__18___redArg(x_92); lean_dec_ref(x_118); -x_32 = x_92; -x_33 = x_93; -x_34 = x_117; -x_35 = x_116; -x_36 = x_96; -x_37 = x_115; -x_38 = x_99; -x_39 = x_91; -x_40 = x_101; -x_41 = x_100; +x_32 = x_88; +x_33 = x_115; +x_34 = x_90; +x_35 = x_91; +x_36 = x_92; +x_37 = x_117; +x_38 = x_116; +x_39 = x_97; +x_40 = x_98; +x_41 = x_101; x_42 = lean_box(0); goto block_85; } else { -x_32 = x_92; -x_33 = x_93; -x_34 = x_117; -x_35 = x_116; -x_36 = x_96; -x_37 = x_115; -x_38 = x_99; -x_39 = x_91; -x_40 = x_101; -x_41 = x_100; +x_32 = x_88; +x_33 = x_115; +x_34 = x_90; +x_35 = x_91; +x_36 = x_92; +x_37 = x_117; +x_38 = x_116; +x_39 = x_97; +x_40 = x_98; +x_41 = x_101; x_42 = lean_box(0); goto block_85; } @@ -9114,15 +9114,15 @@ lean_dec(x_110); lean_dec(x_108); lean_dec_ref(x_102); lean_dec(x_101); -lean_dec_ref(x_100); -lean_dec_ref(x_99); -lean_dec(x_96); +lean_dec(x_100); +lean_dec_ref(x_98); +lean_dec_ref(x_97); lean_dec(x_95); -lean_dec(x_94); -lean_dec_ref(x_93); lean_dec(x_92); lean_dec_ref(x_91); +lean_dec_ref(x_90); lean_dec(x_89); +lean_dec(x_88); x_119 = !lean_is_exclusive(x_111); if (x_119 == 0) { @@ -9146,15 +9146,15 @@ uint8_t x_122; lean_dec(x_108); lean_dec_ref(x_102); lean_dec(x_101); -lean_dec_ref(x_100); -lean_dec_ref(x_99); -lean_dec(x_96); +lean_dec(x_100); +lean_dec_ref(x_98); +lean_dec_ref(x_97); lean_dec(x_95); -lean_dec(x_94); -lean_dec_ref(x_93); lean_dec(x_92); lean_dec_ref(x_91); +lean_dec_ref(x_90); lean_dec(x_89); +lean_dec(x_88); x_122 = !lean_is_exclusive(x_109); if (x_122 == 0) { @@ -9177,15 +9177,15 @@ else uint8_t x_125; lean_dec_ref(x_102); lean_dec(x_101); -lean_dec_ref(x_100); -lean_dec_ref(x_99); -lean_dec(x_96); +lean_dec(x_100); +lean_dec_ref(x_98); +lean_dec_ref(x_97); lean_dec(x_95); -lean_dec(x_94); -lean_dec_ref(x_93); lean_dec(x_92); lean_dec_ref(x_91); +lean_dec_ref(x_90); lean_dec(x_89); +lean_dec(x_88); x_125 = !lean_is_exclusive(x_107); if (x_125 == 0) { @@ -9206,105 +9206,105 @@ return x_127; block_195: { 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; size_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; 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_inc_ref(x_134); -x_150 = l_Array_append___redArg(x_134, x_149); +lean_inc_ref(x_144); +x_150 = l_Array_append___redArg(x_144, x_149); lean_dec_ref(x_149); -lean_inc(x_137); -lean_inc(x_142); +lean_inc(x_129); +lean_inc(x_133); x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_142); -lean_ctor_set(x_151, 1, x_137); +lean_ctor_set(x_151, 0, x_133); +lean_ctor_set(x_151, 1, x_129); lean_ctor_set(x_151, 2, x_150); x_152 = l_Lean_Elab_Command_elabNotation___closed__8; x_153 = l_Lean_Elab_Command_mkUnexpander___closed__54; -lean_inc(x_142); +lean_inc(x_133); x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_142); +lean_ctor_set(x_154, 0, x_133); lean_ctor_set(x_154, 1, x_153); x_155 = l_Lean_Elab_Command_elabNotation___closed__9; -lean_inc(x_142); +lean_inc(x_133); x_156 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_156, 0, x_142); +lean_ctor_set(x_156, 0, x_133); lean_ctor_set(x_156, 1, x_155); x_157 = l_Lean_Elab_Command_mkUnexpander___closed__27; -lean_inc(x_142); +lean_inc(x_133); x_158 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_158, 0, x_142); +lean_ctor_set(x_158, 0, x_133); lean_ctor_set(x_158, 1, x_157); -x_159 = l_Nat_reprFast(x_133); +x_159 = l_Nat_reprFast(x_131); x_160 = lean_box(2); x_161 = l_Lean_Syntax_mkNumLit(x_159, x_160); x_162 = l_Lean_Elab_Command_mkUnexpander___closed__38; -lean_inc(x_142); +lean_inc(x_133); x_163 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_163, 0, x_142); +lean_ctor_set(x_163, 0, x_133); lean_ctor_set(x_163, 1, x_162); -lean_inc(x_142); -x_164 = l_Lean_Syntax_node5(x_142, x_152, x_154, x_156, x_158, x_161, x_163); -lean_inc(x_137); -lean_inc(x_142); -x_165 = l_Lean_Syntax_node1(x_142, x_137, x_164); -x_166 = lean_array_size(x_140); -x_167 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabNotation_spec__21(x_166, x_130, x_140); -lean_inc_ref(x_134); -x_168 = l_Array_append___redArg(x_134, x_167); +lean_inc(x_133); +x_164 = l_Lean_Syntax_node5(x_133, x_152, x_154, x_156, x_158, x_161, x_163); +lean_inc(x_129); +lean_inc(x_133); +x_165 = l_Lean_Syntax_node1(x_133, x_129, x_164); +x_166 = lean_array_size(x_141); +x_167 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_elabNotation_spec__21(x_166, x_136, x_141); +lean_inc_ref(x_144); +x_168 = l_Array_append___redArg(x_144, x_167); lean_dec_ref(x_167); -lean_inc(x_137); -lean_inc(x_142); +lean_inc(x_129); +lean_inc(x_133); x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_142); -lean_ctor_set(x_169, 1, x_137); +lean_ctor_set(x_169, 0, x_133); +lean_ctor_set(x_169, 1, x_129); lean_ctor_set(x_169, 2, x_168); x_170 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__7; -lean_inc(x_142); +lean_inc(x_133); x_171 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_171, 0, x_142); +lean_ctor_set(x_171, 0, x_133); lean_ctor_set(x_171, 1, x_170); x_172 = l_Lean_Elab_Command_elabNotation___closed__10; x_173 = lean_array_push(x_172, x_143); -x_174 = lean_array_push(x_173, x_145); -lean_inc(x_148); -x_175 = lean_array_push(x_174, x_148); -x_176 = lean_array_push(x_175, x_136); -x_177 = lean_array_push(x_176, x_131); +x_174 = lean_array_push(x_173, x_148); +lean_inc(x_147); +x_175 = lean_array_push(x_174, x_147); +x_176 = lean_array_push(x_175, x_137); +x_177 = lean_array_push(x_176, x_140); x_178 = lean_array_push(x_177, x_151); x_179 = lean_array_push(x_178, x_165); x_180 = lean_array_push(x_179, x_169); x_181 = lean_array_push(x_180, x_171); -x_182 = lean_array_push(x_181, x_141); +x_182 = lean_array_push(x_181, x_146); x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_142); -lean_ctor_set(x_183, 1, x_138); +lean_ctor_set(x_183, 0, x_133); +lean_ctor_set(x_183, 1, x_139); lean_ctor_set(x_183, 2, x_182); -lean_inc(x_135); -lean_inc_ref(x_146); -x_184 = l_Lean_Elab_Command_elabSyntax(x_183, x_146, x_135); +lean_inc(x_132); +lean_inc_ref(x_145); +x_184 = l_Lean_Elab_Command_elabSyntax(x_183, x_145, x_132); if (lean_obj_tag(x_184) == 0) { lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; x_185 = lean_ctor_get(x_184, 0); lean_inc(x_185); lean_dec_ref(x_184); -x_186 = lean_array_get_size(x_144); +x_186 = lean_array_get_size(x_134); x_187 = l_Lean_Elab_Command_mkUnexpander___closed__70; x_188 = lean_nat_dec_lt(x_87, x_186); if (x_188 == 0) { lean_dec(x_186); -x_88 = x_130; -x_89 = x_132; -x_90 = lean_box(0); -x_91 = x_134; -x_92 = x_135; -x_93 = x_162; -x_94 = x_185; -x_95 = x_160; -x_96 = x_137; -x_97 = x_139; -x_98 = x_144; -x_99 = x_146; -x_100 = x_147; -x_101 = x_148; +x_88 = x_129; +x_89 = x_160; +x_90 = x_162; +x_91 = x_130; +x_92 = x_132; +x_93 = x_134; +x_94 = x_136; +x_95 = x_138; +x_96 = x_142; +x_97 = x_144; +x_98 = x_145; +x_99 = lean_box(0); +x_100 = x_185; +x_101 = x_147; x_102 = x_187; goto block_128; } @@ -9315,20 +9315,20 @@ x_189 = lean_nat_dec_le(x_186, x_186); if (x_189 == 0) { lean_dec(x_186); -x_88 = x_130; -x_89 = x_132; -x_90 = lean_box(0); -x_91 = x_134; -x_92 = x_135; -x_93 = x_162; -x_94 = x_185; -x_95 = x_160; -x_96 = x_137; -x_97 = x_139; -x_98 = x_144; -x_99 = x_146; -x_100 = x_147; -x_101 = x_148; +x_88 = x_129; +x_89 = x_160; +x_90 = x_162; +x_91 = x_130; +x_92 = x_132; +x_93 = x_134; +x_94 = x_136; +x_95 = x_138; +x_96 = x_142; +x_97 = x_144; +x_98 = x_145; +x_99 = lean_box(0); +x_100 = x_185; +x_101 = x_147; x_102 = x_187; goto block_128; } @@ -9337,21 +9337,21 @@ else size_t x_190; lean_object* x_191; x_190 = lean_usize_of_nat(x_186); lean_dec(x_186); -x_191 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Command_elabNotation_spec__25(x_27, x_28, x_29, x_144, x_130, x_190, x_187); -x_88 = x_130; -x_89 = x_132; -x_90 = lean_box(0); -x_91 = x_134; -x_92 = x_135; -x_93 = x_162; -x_94 = x_185; -x_95 = x_160; -x_96 = x_137; -x_97 = x_139; -x_98 = x_144; -x_99 = x_146; -x_100 = x_147; -x_101 = x_148; +x_191 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Command_elabNotation_spec__25(x_27, x_28, x_29, x_134, x_136, x_190, x_187); +x_88 = x_129; +x_89 = x_160; +x_90 = x_162; +x_91 = x_130; +x_92 = x_132; +x_93 = x_134; +x_94 = x_136; +x_95 = x_138; +x_96 = x_142; +x_97 = x_144; +x_98 = x_145; +x_99 = lean_box(0); +x_100 = x_185; +x_101 = x_147; x_102 = x_191; goto block_128; } @@ -9360,14 +9360,14 @@ goto block_128; else { uint8_t x_192; -lean_dec(x_148); -lean_dec_ref(x_147); -lean_dec_ref(x_146); +lean_dec(x_147); +lean_dec_ref(x_145); lean_dec_ref(x_144); -lean_dec(x_137); -lean_dec(x_135); +lean_dec(x_138); lean_dec_ref(x_134); lean_dec(x_132); +lean_dec_ref(x_130); +lean_dec(x_129); x_192 = !lean_is_exclusive(x_184); if (x_192 == 0) { @@ -9388,92 +9388,92 @@ return x_194; block_232: { lean_object* x_217; lean_object* x_218; -lean_inc_ref(x_201); -x_217 = l_Array_append___redArg(x_201, x_216); +lean_inc_ref(x_211); +x_217 = l_Array_append___redArg(x_211, x_216); lean_dec_ref(x_216); -lean_inc(x_204); -lean_inc(x_210); +lean_inc(x_196); +lean_inc(x_199); x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_210); -lean_ctor_set(x_218, 1, x_204); +lean_ctor_set(x_218, 0, x_199); +lean_ctor_set(x_218, 1, x_196); lean_ctor_set(x_218, 2, x_217); -if (lean_obj_tag(x_198) == 0) +if (lean_obj_tag(x_208) == 0) { lean_object* x_219; x_219 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_129 = lean_box(0); +x_129 = x_196; x_130 = x_197; -x_131 = x_218; -x_132 = x_199; -x_133 = x_200; +x_131 = x_198; +x_132 = x_200; +x_133 = x_199; x_134 = x_201; -x_135 = x_202; +x_135 = lean_box(0); x_136 = x_203; x_137 = x_204; x_138 = x_205; x_139 = x_206; -x_140 = x_207; -x_141 = x_208; +x_140 = x_218; +x_141 = x_207; x_142 = x_210; x_143 = x_209; x_144 = x_211; -x_145 = x_212; -x_146 = x_213; -x_147 = x_215; -x_148 = x_214; +x_145 = x_213; +x_146 = x_212; +x_147 = x_214; +x_148 = x_215; x_149 = x_219; goto block_195; } else { 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_220 = lean_ctor_get(x_198, 0); +x_220 = lean_ctor_get(x_208, 0); lean_inc(x_220); -lean_dec_ref(x_198); +lean_dec_ref(x_208); x_221 = l_Lean_Elab_Command_elabNotation___closed__12; x_222 = l_Lean_Elab_Command_mkUnexpander___closed__54; -lean_inc(x_210); +lean_inc(x_199); x_223 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_223, 0, x_210); +lean_ctor_set(x_223, 0, x_199); lean_ctor_set(x_223, 1, x_222); x_224 = l_Lean_Elab_Command_elabNotation___closed__13; -lean_inc(x_210); +lean_inc(x_199); x_225 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_225, 0, x_210); +lean_ctor_set(x_225, 0, x_199); lean_ctor_set(x_225, 1, x_224); x_226 = l_Lean_Elab_Command_mkUnexpander___closed__27; -lean_inc(x_210); +lean_inc(x_199); x_227 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_227, 0, x_210); +lean_ctor_set(x_227, 0, x_199); lean_ctor_set(x_227, 1, x_226); x_228 = l_Lean_Elab_Command_mkUnexpander___closed__38; -lean_inc(x_210); +lean_inc(x_199); x_229 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_229, 0, x_210); +lean_ctor_set(x_229, 0, x_199); lean_ctor_set(x_229, 1, x_228); -lean_inc(x_210); -x_230 = l_Lean_Syntax_node5(x_210, x_221, x_223, x_225, x_227, x_220, x_229); +lean_inc(x_199); +x_230 = l_Lean_Syntax_node5(x_199, x_221, x_223, x_225, x_227, x_220, x_229); x_231 = l_Array_mkArray1___redArg(x_230); -x_129 = lean_box(0); +x_129 = x_196; x_130 = x_197; -x_131 = x_218; -x_132 = x_199; -x_133 = x_200; +x_131 = x_198; +x_132 = x_200; +x_133 = x_199; x_134 = x_201; -x_135 = x_202; +x_135 = lean_box(0); x_136 = x_203; x_137 = x_204; x_138 = x_205; x_139 = x_206; -x_140 = x_207; -x_141 = x_208; +x_140 = x_218; +x_141 = x_207; x_142 = x_210; x_143 = x_209; x_144 = x_211; -x_145 = x_212; -x_146 = x_213; -x_147 = x_215; -x_148 = x_214; +x_145 = x_213; +x_146 = x_212; +x_147 = x_214; +x_148 = x_215; x_149 = x_231; goto block_195; } @@ -9481,81 +9481,81 @@ goto block_195; block_264: { lean_object* x_254; lean_object* x_255; lean_object* x_256; -lean_inc_ref(x_239); -x_254 = l_Array_append___redArg(x_239, x_253); +lean_inc_ref(x_248); +x_254 = l_Array_append___redArg(x_248, x_253); lean_dec_ref(x_253); -lean_inc(x_241); -lean_inc(x_248); +lean_inc(x_233); +lean_inc(x_236); x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_248); -lean_ctor_set(x_255, 1, x_241); +lean_ctor_set(x_255, 0, x_236); +lean_ctor_set(x_255, 1, x_233); lean_ctor_set(x_255, 2, x_254); -lean_inc(x_248); +lean_inc(x_236); x_256 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_256, 0, x_248); -lean_ctor_set(x_256, 1, x_235); -if (lean_obj_tag(x_246) == 0) +lean_ctor_set(x_256, 0, x_236); +lean_ctor_set(x_256, 1, x_245); +if (lean_obj_tag(x_251) == 0) { lean_object* x_257; x_257 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_196 = lean_box(0); +x_196 = x_233; x_197 = x_234; -x_198 = x_236; -x_199 = x_237; -x_200 = x_238; -x_201 = x_239; -x_202 = x_240; -x_203 = x_256; -x_204 = x_241; -x_205 = x_242; -x_206 = x_244; -x_207 = x_245; -x_208 = x_243; +x_198 = x_235; +x_199 = x_236; +x_200 = x_237; +x_201 = x_238; +x_202 = lean_box(0); +x_203 = x_240; +x_204 = x_256; +x_205 = x_241; +x_206 = x_242; +x_207 = x_243; +x_208 = x_244; x_209 = x_247; -x_210 = x_248; -x_211 = x_249; -x_212 = x_255; -x_213 = x_250; +x_210 = x_246; +x_211 = x_248; +x_212 = x_250; +x_213 = x_249; x_214 = x_252; -x_215 = x_251; +x_215 = x_255; x_216 = x_257; goto block_232; } else { lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -x_258 = lean_ctor_get(x_246, 0); +x_258 = lean_ctor_get(x_251, 0); lean_inc(x_258); -lean_dec_ref(x_246); +lean_dec_ref(x_251); x_259 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; x_260 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__7; -lean_inc(x_248); +lean_inc(x_236); x_261 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_261, 0, x_248); +lean_ctor_set(x_261, 0, x_236); lean_ctor_set(x_261, 1, x_260); -lean_inc(x_248); -x_262 = l_Lean_Syntax_node2(x_248, x_259, x_261, x_258); +lean_inc(x_236); +x_262 = l_Lean_Syntax_node2(x_236, x_259, x_261, x_258); x_263 = l_Array_mkArray1___redArg(x_262); -x_196 = lean_box(0); +x_196 = x_233; x_197 = x_234; -x_198 = x_236; -x_199 = x_237; -x_200 = x_238; -x_201 = x_239; -x_202 = x_240; -x_203 = x_256; -x_204 = x_241; -x_205 = x_242; -x_206 = x_244; -x_207 = x_245; -x_208 = x_243; +x_198 = x_235; +x_199 = x_236; +x_200 = x_237; +x_201 = x_238; +x_202 = lean_box(0); +x_203 = x_240; +x_204 = x_256; +x_205 = x_241; +x_206 = x_242; +x_207 = x_243; +x_208 = x_244; x_209 = x_247; -x_210 = x_248; -x_211 = x_249; -x_212 = x_255; -x_213 = x_250; +x_210 = x_246; +x_211 = x_248; +x_212 = x_250; +x_213 = x_249; x_214 = x_252; -x_215 = x_251; +x_215 = x_255; x_216 = x_263; goto block_232; } @@ -9563,93 +9563,93 @@ goto block_232; block_300: { lean_object* x_286; lean_object* x_287; -lean_inc_ref(x_271); -x_286 = l_Array_append___redArg(x_271, x_285); +lean_inc_ref(x_280); +x_286 = l_Array_append___redArg(x_280, x_285); lean_dec_ref(x_285); -lean_inc(x_273); -lean_inc(x_280); +lean_inc(x_265); +lean_inc(x_268); x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_280); -lean_ctor_set(x_287, 1, x_273); +lean_ctor_set(x_287, 0, x_268); +lean_ctor_set(x_287, 1, x_265); lean_ctor_set(x_287, 2, x_286); -if (lean_obj_tag(x_279) == 0) +if (lean_obj_tag(x_278) == 0) { lean_object* x_288; x_288 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_233 = lean_box(0); +x_233 = x_265; x_234 = x_266; x_235 = x_267; x_236 = x_268; x_237 = x_269; x_238 = x_270; -x_239 = x_271; +x_239 = lean_box(0); x_240 = x_272; x_241 = x_273; x_242 = x_274; -x_243 = x_276; -x_244 = x_277; -x_245 = x_275; -x_246 = x_278; +x_243 = x_275; +x_244 = x_276; +x_245 = x_277; +x_246 = x_279; x_247 = x_287; x_248 = x_280; -x_249 = x_281; -x_250 = x_282; -x_251 = x_284; -x_252 = x_283; +x_249 = x_282; +x_250 = x_281; +x_251 = x_283; +x_252 = x_284; x_253 = x_288; goto block_264; } else { lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; -x_289 = lean_ctor_get(x_279, 0); +x_289 = lean_ctor_get(x_278, 0); lean_inc(x_289); -lean_dec_ref(x_279); +lean_dec_ref(x_278); x_290 = l_Lean_Elab_Command_mkUnexpander___closed__11; -lean_inc_ref(x_284); -x_291 = l_Lean_Name_mkStr4(x_27, x_28, x_284, x_290); +lean_inc_ref(x_266); +x_291 = l_Lean_Name_mkStr4(x_27, x_28, x_266, x_290); x_292 = l_Lean_Elab_Command_mkUnexpander___closed__13; -lean_inc(x_280); +lean_inc(x_268); x_293 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_293, 0, x_280); +lean_ctor_set(x_293, 0, x_268); lean_ctor_set(x_293, 1, x_292); -lean_inc_ref(x_271); -x_294 = l_Array_append___redArg(x_271, x_289); +lean_inc_ref(x_280); +x_294 = l_Array_append___redArg(x_280, x_289); lean_dec(x_289); -lean_inc(x_273); -lean_inc(x_280); +lean_inc(x_265); +lean_inc(x_268); x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_280); -lean_ctor_set(x_295, 1, x_273); +lean_ctor_set(x_295, 0, x_268); +lean_ctor_set(x_295, 1, x_265); lean_ctor_set(x_295, 2, x_294); x_296 = l_Lean_Elab_Command_mkUnexpander___closed__18; -lean_inc(x_280); +lean_inc(x_268); x_297 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_297, 0, x_280); +lean_ctor_set(x_297, 0, x_268); lean_ctor_set(x_297, 1, x_296); -lean_inc(x_280); -x_298 = l_Lean_Syntax_node3(x_280, x_291, x_293, x_295, x_297); +lean_inc(x_268); +x_298 = l_Lean_Syntax_node3(x_268, x_291, x_293, x_295, x_297); x_299 = l_Array_mkArray1___redArg(x_298); -x_233 = lean_box(0); +x_233 = x_265; x_234 = x_266; x_235 = x_267; x_236 = x_268; x_237 = x_269; x_238 = x_270; -x_239 = x_271; +x_239 = lean_box(0); x_240 = x_272; x_241 = x_273; x_242 = x_274; -x_243 = x_276; -x_244 = x_277; -x_245 = x_275; -x_246 = x_278; +x_243 = x_275; +x_244 = x_276; +x_245 = x_277; +x_246 = x_279; x_247 = x_287; x_248 = x_280; -x_249 = x_281; -x_250 = x_282; -x_251 = x_284; -x_252 = x_283; +x_249 = x_282; +x_250 = x_281; +x_251 = x_283; +x_252 = x_284; x_253 = x_299; goto block_264; } @@ -9661,60 +9661,60 @@ x_318 = l_Lean_Elab_Command_elabNotation___closed__14; x_319 = l_Lean_Elab_Command_elabNotation___closed__15; x_320 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_addInheritDocDefault_spec__0___redArg___closed__6; x_321 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Command_addInheritDocDefault_spec__0___redArg___closed__7; -if (lean_obj_tag(x_302) == 0) +if (lean_obj_tag(x_301) == 0) { lean_object* x_322; x_322 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_265 = lean_box(0); -x_266 = x_301; -x_267 = x_318; -x_268 = x_303; -x_269 = x_304; -x_270 = x_305; -x_271 = x_321; -x_272 = x_306; -x_273 = x_320; +x_265 = x_320; +x_266 = x_302; +x_267 = x_303; +x_268 = x_304; +x_269 = x_305; +x_270 = x_306; +x_271 = lean_box(0); +x_272 = x_307; +x_273 = x_308; x_274 = x_319; -x_275 = x_307; -x_276 = x_308; -x_277 = x_309; -x_278 = x_310; -x_279 = x_311; -x_280 = x_312; +x_275 = x_309; +x_276 = x_310; +x_277 = x_318; +x_278 = x_311; +x_279 = x_312; +x_280 = x_321; x_281 = x_313; x_282 = x_314; -x_283 = x_316; -x_284 = x_315; +x_283 = x_315; +x_284 = x_316; x_285 = x_322; goto block_300; } else { lean_object* x_323; lean_object* x_324; -x_323 = lean_ctor_get(x_302, 0); +x_323 = lean_ctor_get(x_301, 0); lean_inc(x_323); -lean_dec_ref(x_302); +lean_dec_ref(x_301); x_324 = l_Array_mkArray1___redArg(x_323); -x_265 = lean_box(0); -x_266 = x_301; -x_267 = x_318; -x_268 = x_303; -x_269 = x_304; -x_270 = x_305; -x_271 = x_321; -x_272 = x_306; -x_273 = x_320; +x_265 = x_320; +x_266 = x_302; +x_267 = x_303; +x_268 = x_304; +x_269 = x_305; +x_270 = x_306; +x_271 = lean_box(0); +x_272 = x_307; +x_273 = x_308; x_274 = x_319; -x_275 = x_307; -x_276 = x_308; -x_277 = x_309; -x_278 = x_310; -x_279 = x_311; -x_280 = x_312; +x_275 = x_309; +x_276 = x_310; +x_277 = x_318; +x_278 = x_311; +x_279 = x_312; +x_280 = x_321; x_281 = x_313; x_282 = x_314; -x_283 = x_316; -x_284 = x_315; +x_283 = x_315; +x_284 = x_316; x_285 = x_324; goto block_300; } @@ -9781,43 +9781,43 @@ if (lean_obj_tag(x_352) == 0) lean_object* x_360; x_360 = l_Lean_getMainModule___at___00Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabNotation_spec__1_spec__18___redArg(x_334); lean_dec_ref(x_360); -x_301 = x_343; -x_302 = x_326; -x_303 = x_327; -x_304 = x_354; -x_305 = x_338; -x_306 = x_334; -x_307 = x_348; -x_308 = x_358; -x_309 = x_357; -x_310 = x_328; +x_301 = x_328; +x_302 = x_327; +x_303 = x_338; +x_304 = x_359; +x_305 = x_334; +x_306 = x_341; +x_307 = x_343; +x_308 = x_354; +x_309 = x_348; +x_310 = x_326; x_311 = x_356; -x_312 = x_359; -x_313 = x_341; +x_312 = x_357; +x_313 = x_358; x_314 = x_333; -x_315 = x_331; -x_316 = x_330; +x_315 = x_330; +x_316 = x_331; x_317 = lean_box(0); goto block_325; } else { -x_301 = x_343; -x_302 = x_326; -x_303 = x_327; -x_304 = x_354; -x_305 = x_338; -x_306 = x_334; -x_307 = x_348; -x_308 = x_358; -x_309 = x_357; -x_310 = x_328; +x_301 = x_328; +x_302 = x_327; +x_303 = x_338; +x_304 = x_359; +x_305 = x_334; +x_306 = x_341; +x_307 = x_343; +x_308 = x_354; +x_309 = x_348; +x_310 = x_326; x_311 = x_356; -x_312 = x_359; -x_313 = x_341; +x_312 = x_357; +x_313 = x_358; x_314 = x_333; -x_315 = x_331; -x_316 = x_330; +x_315 = x_330; +x_316 = x_331; x_317 = lean_box(0); goto block_325; } @@ -9831,11 +9831,11 @@ lean_dec_ref(x_341); lean_dec(x_338); lean_dec(x_334); lean_dec_ref(x_333); -lean_dec_ref(x_331); +lean_dec(x_331); lean_dec(x_330); lean_dec(x_329); lean_dec(x_328); -lean_dec(x_327); +lean_dec_ref(x_327); lean_dec(x_326); lean_dec(x_1); x_361 = !lean_is_exclusive(x_351); @@ -9863,11 +9863,11 @@ lean_dec_ref(x_341); lean_dec(x_338); lean_dec(x_334); lean_dec_ref(x_333); -lean_dec_ref(x_331); +lean_dec(x_331); lean_dec(x_330); lean_dec(x_329); lean_dec(x_328); -lean_dec(x_327); +lean_dec_ref(x_327); lean_dec(x_326); lean_dec(x_1); x_364 = !lean_is_exclusive(x_349); @@ -9894,11 +9894,11 @@ lean_dec_ref(x_341); lean_dec(x_338); lean_dec(x_334); lean_dec_ref(x_333); -lean_dec_ref(x_331); +lean_dec(x_331); lean_dec(x_330); lean_dec(x_329); lean_dec(x_328); -lean_dec(x_327); +lean_dec_ref(x_327); lean_dec(x_326); lean_dec(x_1); x_367 = !lean_is_exclusive(x_347); @@ -9923,11 +9923,11 @@ else uint8_t x_370; lean_dec(x_334); lean_dec_ref(x_333); -lean_dec_ref(x_331); +lean_dec(x_331); lean_dec(x_330); lean_dec(x_329); lean_dec(x_328); -lean_dec(x_327); +lean_dec_ref(x_327); lean_dec(x_326); lean_dec(x_1); x_370 = !lean_is_exclusive(x_337); @@ -9957,7 +9957,7 @@ if (x_387 == 0) { uint8_t x_388; lean_inc(x_386); -x_388 = l_Lean_Syntax_matchesNull(x_386, x_378); +x_388 = l_Lean_Syntax_matchesNull(x_386, x_374); if (x_388 == 0) { lean_object* x_389; @@ -9966,10 +9966,10 @@ lean_dec(x_383); lean_dec_ref(x_382); lean_dec(x_381); lean_dec(x_380); -lean_dec_ref(x_379); -lean_dec(x_377); +lean_dec(x_378); +lean_dec_ref(x_377); lean_dec(x_376); -lean_dec(x_374); +lean_dec(x_375); lean_dec(x_1); x_389 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_389; @@ -9990,10 +9990,10 @@ lean_dec(x_383); lean_dec_ref(x_382); lean_dec(x_381); lean_dec(x_380); -lean_dec_ref(x_379); -lean_dec(x_377); +lean_dec(x_378); +lean_dec_ref(x_377); lean_dec(x_376); -lean_dec(x_374); +lean_dec(x_375); lean_dec(x_1); x_393 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_393; @@ -10001,16 +10001,16 @@ return x_393; else { lean_object* x_394; lean_object* x_395; -x_394 = l_Lean_Syntax_getArg(x_390, x_375); +x_394 = l_Lean_Syntax_getArg(x_390, x_379); lean_dec(x_390); x_395 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_395, 0, x_394); -x_326 = x_374; -x_327 = x_381; +x_326 = x_381; +x_327 = x_377; x_328 = x_376; -x_329 = x_377; -x_330 = x_380; -x_331 = x_379; +x_329 = x_375; +x_330 = x_378; +x_331 = x_380; x_332 = x_395; x_333 = x_382; x_334 = x_383; @@ -10024,12 +10024,12 @@ else lean_object* x_396; lean_dec(x_386); x_396 = lean_box(0); -x_326 = x_374; -x_327 = x_381; +x_326 = x_381; +x_327 = x_377; x_328 = x_376; -x_329 = x_377; -x_330 = x_380; -x_331 = x_379; +x_329 = x_375; +x_330 = x_378; +x_331 = x_380; x_332 = x_396; x_333 = x_382; x_334 = x_383; @@ -10047,7 +10047,7 @@ if (x_410 == 0) { uint8_t x_411; lean_inc(x_409); -x_411 = l_Lean_Syntax_matchesNull(x_409, x_403); +x_411 = l_Lean_Syntax_matchesNull(x_409, x_398); if (x_411 == 0) { lean_object* x_412; @@ -10055,10 +10055,10 @@ lean_dec(x_409); lean_dec(x_406); lean_dec_ref(x_405); lean_dec(x_404); -lean_dec_ref(x_402); +lean_dec(x_402); lean_dec(x_401); lean_dec(x_400); -lean_dec(x_398); +lean_dec_ref(x_399); lean_dec(x_1); x_412 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_412; @@ -10078,10 +10078,10 @@ lean_dec(x_413); lean_dec(x_406); lean_dec_ref(x_405); lean_dec(x_404); -lean_dec_ref(x_402); +lean_dec(x_402); lean_dec(x_401); lean_dec(x_400); -lean_dec(x_398); +lean_dec_ref(x_399); lean_dec(x_1); x_416 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_416; @@ -10089,17 +10089,17 @@ return x_416; else { lean_object* x_417; lean_object* x_418; -x_417 = l_Lean_Syntax_getArg(x_413, x_399); +x_417 = l_Lean_Syntax_getArg(x_413, x_403); lean_dec(x_413); x_418 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_418, 0, x_417); x_374 = x_398; -x_375 = x_399; -x_376 = x_404; -x_377 = x_400; -x_378 = x_403; -x_379 = x_402; -x_380 = x_401; +x_375 = x_401; +x_376 = x_400; +x_377 = x_399; +x_378 = x_404; +x_379 = x_403; +x_380 = x_402; x_381 = x_418; x_382 = x_405; x_383 = x_406; @@ -10114,12 +10114,12 @@ lean_object* x_419; lean_dec(x_409); x_419 = lean_box(0); x_374 = x_398; -x_375 = x_399; -x_376 = x_404; -x_377 = x_400; -x_378 = x_403; -x_379 = x_402; -x_380 = x_401; +x_375 = x_401; +x_376 = x_400; +x_377 = x_399; +x_378 = x_404; +x_379 = x_403; +x_380 = x_402; x_381 = x_419; x_382 = x_405; x_383 = x_406; @@ -10143,7 +10143,7 @@ lean_dec(x_428); lean_dec(x_425); lean_dec_ref(x_424); lean_dec(x_423); -lean_dec(x_421); +lean_dec(x_422); lean_dec(x_1); x_432 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_432; @@ -10159,7 +10159,7 @@ if (x_436 == 0) { uint8_t x_437; lean_inc(x_435); -x_437 = l_Lean_Syntax_matchesNull(x_435, x_422); +x_437 = l_Lean_Syntax_matchesNull(x_435, x_421); if (x_437 == 0) { lean_object* x_438; @@ -10168,7 +10168,7 @@ lean_dec(x_428); lean_dec(x_425); lean_dec_ref(x_424); lean_dec(x_423); -lean_dec(x_421); +lean_dec(x_422); lean_dec(x_1); x_438 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_438; @@ -10189,7 +10189,7 @@ lean_dec(x_428); lean_dec(x_425); lean_dec_ref(x_424); lean_dec(x_423); -lean_dec(x_421); +lean_dec(x_422); lean_dec(x_1); x_442 = l_Lean_Elab_throwUnsupportedSyntax___at___00Lean_Elab_Command_elabNotation_spec__0___redArg(); return x_442; @@ -10197,16 +10197,16 @@ return x_442; else { lean_object* x_443; lean_object* x_444; -x_443 = l_Lean_Syntax_getArg(x_439, x_422); +x_443 = l_Lean_Syntax_getArg(x_439, x_421); lean_dec(x_439); x_444 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_444, 0, x_443); x_398 = x_421; -x_399 = x_433; -x_400 = x_423; -x_401 = x_428; -x_402 = x_429; -x_403 = x_422; +x_399 = x_429; +x_400 = x_422; +x_401 = x_423; +x_402 = x_428; +x_403 = x_433; x_404 = x_444; x_405 = x_424; x_406 = x_425; @@ -10221,11 +10221,11 @@ lean_object* x_445; lean_dec(x_435); x_445 = lean_box(0); x_398 = x_421; -x_399 = x_433; -x_400 = x_423; -x_401 = x_428; -x_402 = x_429; -x_403 = x_422; +x_399 = x_429; +x_400 = x_422; +x_401 = x_423; +x_402 = x_428; +x_403 = x_433; x_404 = x_445; x_405 = x_424; x_406 = x_425; @@ -10284,8 +10284,8 @@ x_461 = l_Lean_Syntax_getArgs(x_460); lean_dec(x_460); x_462 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_462, 0, x_461); -x_421 = x_447; -x_422 = x_451; +x_421 = x_451; +x_422 = x_447; x_423 = x_462; x_424 = x_448; x_425 = x_449; @@ -10299,8 +10299,8 @@ else lean_object* x_463; lean_dec(x_452); x_463 = lean_box(0); -x_421 = x_447; -x_422 = x_451; +x_421 = x_451; +x_422 = x_447; x_423 = x_463; x_424 = x_448; x_425 = x_449; @@ -10314,8 +10314,8 @@ goto block_446; lean_object* x_11; lean_object* x_12; x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Command_mkUnexpander), 5, 3); lean_closure_set(x_11, 0, x_7); -lean_closure_set(x_11, 1, x_5); -lean_closure_set(x_11, 2, x_6); +lean_closure_set(x_11, 1, x_6); +lean_closure_set(x_11, 2, x_5); lean_inc_ref(x_8); x_12 = l_Lean_Elab_liftMacroM___at___00Lean_Elab_Command_elabNotation_spec__1___redArg(x_11, x_8, x_9); if (lean_obj_tag(x_12) == 0) @@ -10400,110 +10400,110 @@ return x_25; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; x_43 = l_Lean_Elab_Command_elabNotation___closed__2; x_44 = l_Lean_Elab_Command_elabNotation___closed__3; -lean_inc(x_36); -lean_inc(x_34); +lean_inc(x_32); +lean_inc(x_37); x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_34); -lean_ctor_set(x_45, 1, x_36); +lean_ctor_set(x_45, 0, x_37); +lean_ctor_set(x_45, 1, x_32); lean_ctor_set(x_45, 2, x_39); -lean_inc(x_34); +lean_inc(x_37); x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_34); +lean_ctor_set(x_46, 0, x_37); lean_ctor_set(x_46, 1, x_43); x_47 = l_Lean_Elab_Command_mkUnexpander___closed__30; -lean_inc_ref(x_41); -x_48 = l_Lean_Name_mkStr4(x_27, x_28, x_41, x_47); +lean_inc_ref(x_35); +x_48 = l_Lean_Name_mkStr4(x_27, x_28, x_35, x_47); x_49 = l_Lean_Elab_Command_mkUnexpander___closed__32; -lean_inc_ref(x_41); -x_50 = l_Lean_Name_mkStr4(x_27, x_28, x_41, x_49); +lean_inc_ref(x_35); +x_50 = l_Lean_Name_mkStr4(x_27, x_28, x_35, x_49); x_51 = l_Lean_Elab_Command_mkUnexpander___closed__34; -lean_inc(x_34); +lean_inc(x_37); x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_34); +lean_ctor_set(x_52, 0, x_37); lean_ctor_set(x_52, 1, x_51); x_53 = l_Lean_Elab_Command_mkUnexpander___closed__35; -lean_inc_ref(x_41); -x_54 = l_Lean_Name_mkStr4(x_27, x_28, x_41, x_53); +lean_inc_ref(x_35); +x_54 = l_Lean_Name_mkStr4(x_27, x_28, x_35, x_53); x_55 = l_Lean_Elab_Command_mkUnexpander___closed__37; -lean_inc(x_34); +lean_inc(x_37); x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_34); +lean_ctor_set(x_56, 0, x_37); lean_ctor_set(x_56, 1, x_55); -lean_inc(x_34); +lean_inc(x_37); x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_34); -lean_ctor_set(x_57, 1, x_33); +lean_ctor_set(x_57, 0, x_37); +lean_ctor_set(x_57, 1, x_34); lean_inc_ref(x_57); -lean_inc(x_35); +lean_inc(x_38); lean_inc_ref(x_56); lean_inc(x_54); -lean_inc(x_34); -x_58 = l_Lean_Syntax_node3(x_34, x_54, x_56, x_35, x_57); -lean_inc(x_36); -lean_inc(x_34); -x_59 = l_Lean_Syntax_node1(x_34, x_36, x_58); -lean_inc(x_36); -lean_inc(x_34); -x_60 = l_Lean_Syntax_node1(x_34, x_36, x_59); +lean_inc(x_37); +x_58 = l_Lean_Syntax_node3(x_37, x_54, x_56, x_38, x_57); +lean_inc(x_32); +lean_inc(x_37); +x_59 = l_Lean_Syntax_node1(x_37, x_32, x_58); +lean_inc(x_32); +lean_inc(x_37); +x_60 = l_Lean_Syntax_node1(x_37, x_32, x_59); x_61 = l_Lean_Elab_Command_mkUnexpander___closed__39; -lean_inc(x_34); +lean_inc(x_37); x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_34); +lean_ctor_set(x_62, 0, x_37); lean_ctor_set(x_62, 1, x_61); x_63 = l_Lean_Elab_Command_elabNotation___closed__4; -x_64 = l_Lean_Name_mkStr4(x_27, x_28, x_41, x_63); +x_64 = l_Lean_Name_mkStr4(x_27, x_28, x_35, x_63); x_65 = l_Lean_Elab_Command_elabNotation___closed__5; -lean_inc(x_34); +lean_inc(x_37); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_34); +lean_ctor_set(x_66, 0, x_37); lean_ctor_set(x_66, 1, x_65); +lean_inc(x_33); lean_inc(x_37); -lean_inc(x_34); -x_67 = l_Lean_Syntax_node3(x_34, x_54, x_56, x_37, x_57); -lean_inc(x_34); -x_68 = l_Lean_Syntax_node2(x_34, x_64, x_66, x_67); -lean_inc(x_34); -x_69 = l_Lean_Syntax_node4(x_34, x_50, x_52, x_60, x_62, x_68); -lean_inc(x_34); -x_70 = l_Lean_Syntax_node1(x_34, x_36, x_69); -lean_inc(x_34); -x_71 = l_Lean_Syntax_node1(x_34, x_48, x_70); -lean_inc(x_40); +x_67 = l_Lean_Syntax_node3(x_37, x_54, x_56, x_33, x_57); +lean_inc(x_37); +x_68 = l_Lean_Syntax_node2(x_37, x_64, x_66, x_67); +lean_inc(x_37); +x_69 = l_Lean_Syntax_node4(x_37, x_50, x_52, x_60, x_62, x_68); +lean_inc(x_37); +x_70 = l_Lean_Syntax_node1(x_37, x_32, x_69); +lean_inc(x_37); +x_71 = l_Lean_Syntax_node1(x_37, x_48, x_70); +lean_inc(x_41); lean_inc_ref_n(x_45, 2); -x_72 = l_Lean_Syntax_node6(x_34, x_44, x_45, x_45, x_40, x_46, x_45, x_71); -lean_inc(x_40); -x_73 = l_Lean_Elab_Command_isLocalAttrKind(x_40); +x_72 = l_Lean_Syntax_node6(x_37, x_44, x_45, x_45, x_41, x_46, x_45, x_71); +lean_inc(x_41); +x_73 = l_Lean_Elab_Command_isLocalAttrKind(x_41); if (x_73 == 0) { lean_object* x_74; -lean_inc(x_32); -lean_inc_ref(x_38); -x_74 = l_Lean_Elab_Command_elabCommand(x_72, x_38, x_32); +lean_inc(x_36); +lean_inc_ref(x_40); +x_74 = l_Lean_Elab_Command_elabCommand(x_72, x_40, x_36); if (lean_obj_tag(x_74) == 0) { lean_dec_ref(x_74); -x_5 = x_35; -x_6 = x_37; -x_7 = x_40; -x_8 = x_38; -x_9 = x_32; +x_5 = x_33; +x_6 = x_38; +x_7 = x_41; +x_8 = x_40; +x_9 = x_36; x_10 = lean_box(0); goto block_26; } else { -lean_dec(x_40); -lean_dec_ref(x_38); -lean_dec(x_37); -lean_dec(x_35); -lean_dec(x_32); +lean_dec(x_41); +lean_dec_ref(x_40); +lean_dec(x_38); +lean_dec(x_36); +lean_dec(x_33); 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; -x_75 = lean_st_ref_get(x_32); +x_75 = lean_st_ref_get(x_36); x_76 = lean_ctor_get(x_75, 2); lean_inc(x_76); lean_dec_ref(x_75); @@ -10519,27 +10519,27 @@ x_82 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabNotation___lam__0), 2, lean_closure_set(x_82, 0, x_81); x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___boxed), 4, 1); lean_closure_set(x_83, 0, x_72); -lean_inc(x_32); -lean_inc_ref(x_38); -x_84 = l_Lean_Elab_Command_withScope___redArg(x_82, x_83, x_38, x_32); +lean_inc(x_36); +lean_inc_ref(x_40); +x_84 = l_Lean_Elab_Command_withScope___redArg(x_82, x_83, x_40, x_36); if (lean_obj_tag(x_84) == 0) { lean_dec_ref(x_84); -x_5 = x_35; -x_6 = x_37; -x_7 = x_40; -x_8 = x_38; -x_9 = x_32; +x_5 = x_33; +x_6 = x_38; +x_7 = x_41; +x_8 = x_40; +x_9 = x_36; x_10 = lean_box(0); goto block_26; } else { -lean_dec(x_40); -lean_dec_ref(x_38); -lean_dec(x_37); -lean_dec(x_35); -lean_dec(x_32); +lean_dec(x_41); +lean_dec_ref(x_40); +lean_dec(x_38); +lean_dec(x_36); +lean_dec(x_33); return x_84; } } diff --git a/stage0/stdlib/Lean/Elab/PatternVar.c b/stage0/stdlib/Lean/Elab/PatternVar.c index c3e3b9a8cb96..d23ae4fbefcf 100644 --- a/stage0/stdlib/Lean/Elab/PatternVar.c +++ b/stage0/stdlib/Lean/Elab/PatternVar.c @@ -15613,13 +15613,13 @@ goto block_266; block_183: { lean_object* x_162; size_t x_163; size_t x_164; lean_object* x_165; -lean_dec(x_148); +lean_dec(x_152); x_162 = l_Lean_Syntax_TSepArray_getElems___redArg(x_145); lean_dec_ref(x_145); x_163 = lean_array_size(x_162); x_164 = 0; lean_inc_ref(x_159); -x_165 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_CollectPatternVars_collect_spec__0___redArg(x_153, x_141, x_37, x_38, x_39, x_142, x_144, x_151, x_77, x_163, x_164, x_162, x_154, x_155, x_156, x_157, x_158, x_159, x_160); +x_165 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_CollectPatternVars_collect_spec__0___redArg(x_151, x_141, x_37, x_38, x_39, x_142, x_144, x_150, x_77, x_163, x_164, x_162, x_154, x_155, x_156, x_157, x_158, x_159, x_160); if (lean_obj_tag(x_165) == 0) { lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; @@ -15638,30 +15638,30 @@ lean_ctor_set(x_170, 0, x_168); lean_ctor_set(x_170, 1, x_169); x_171 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_CollectPatternVars_collect_spec__0___redArg___closed__3; x_172 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_CollectPatternVars_collect_spec__0___redArg___closed__4; -if (lean_obj_tag(x_143) == 0) +if (lean_obj_tag(x_146) == 0) { lean_object* x_173; x_173 = l_Lean_Elab_Term_CollectPatternVars_instInhabitedState_default___closed__1; -x_114 = x_168; -x_115 = x_166; -x_116 = lean_box(0); -x_117 = x_172; -x_118 = x_146; -x_119 = x_147; +x_114 = x_171; +x_115 = x_143; +x_116 = x_170; +x_117 = x_147; +x_118 = x_148; +x_119 = x_172; x_120 = x_149; -x_121 = x_170; -x_122 = x_150; -x_123 = x_171; -x_124 = x_152; +x_121 = x_166; +x_122 = x_168; +x_123 = lean_box(0); +x_124 = x_153; x_125 = x_173; goto block_138; } 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; -x_174 = lean_ctor_get(x_143, 0); +x_174 = lean_ctor_get(x_146, 0); lean_inc(x_174); -lean_dec_ref(x_143); +lean_dec_ref(x_146); x_175 = l_Array_append___redArg(x_172, x_174); lean_dec(x_174); lean_inc(x_168); @@ -15675,17 +15675,17 @@ x_178 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_178, 0, x_168); lean_ctor_set(x_178, 1, x_177); x_179 = l_Array_mkArray2___redArg(x_176, x_178); -x_114 = x_168; -x_115 = x_166; -x_116 = lean_box(0); -x_117 = x_172; -x_118 = x_146; -x_119 = x_147; +x_114 = x_171; +x_115 = x_143; +x_116 = x_170; +x_117 = x_147; +x_118 = x_148; +x_119 = x_172; x_120 = x_149; -x_121 = x_170; -x_122 = x_150; -x_123 = x_171; -x_124 = x_152; +x_121 = x_166; +x_122 = x_168; +x_123 = lean_box(0); +x_124 = x_153; x_125 = x_179; goto block_138; } @@ -15694,9 +15694,9 @@ else { uint8_t x_180; lean_dec_ref(x_159); -lean_dec_ref(x_152); -lean_dec(x_150); +lean_dec(x_153); lean_dec(x_149); +lean_dec_ref(x_148); lean_dec(x_147); lean_dec(x_146); lean_dec(x_143); @@ -15723,21 +15723,21 @@ lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; x_200 = l_Lean_Elab_Term_CollectPatternVars_collect___lam__0___closed__38; x_201 = l_Lean_Elab_Term_CollectPatternVars_collect___lam__0___closed__39; x_202 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___redArg___closed__3; -x_203 = l_Lean_Syntax_getArgs(x_187); -lean_dec(x_187); -if (lean_obj_tag(x_184) == 0) +x_203 = l_Lean_Syntax_getArgs(x_189); +lean_dec(x_189); +if (lean_obj_tag(x_185) == 0) { -x_143 = x_184; -x_144 = x_185; +x_143 = x_191; +x_144 = x_184; x_145 = x_203; -x_146 = x_191; +x_146 = x_185; x_147 = x_186; -x_148 = x_201; -x_149 = x_188; -x_150 = x_190; -x_151 = x_189; -x_152 = x_202; -x_153 = x_200; +x_148 = x_202; +x_149 = x_187; +x_150 = x_188; +x_151 = x_200; +x_152 = x_201; +x_153 = x_190; x_154 = x_192; x_155 = x_193; x_156 = x_194; @@ -15757,11 +15757,11 @@ lean_dec_ref(x_193); lean_dec(x_192); lean_dec(x_191); lean_dec(x_190); -lean_dec(x_188); +lean_dec(x_187); lean_dec(x_186); -x_204 = lean_ctor_get(x_184, 0); +x_204 = lean_ctor_get(x_185, 0); lean_inc(x_204); -lean_dec_ref(x_184); +lean_dec_ref(x_185); x_205 = l_Lean_Syntax_TSepArray_getElems___redArg(x_204); lean_dec(x_204); x_206 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_CollectPatternVars_collect_spec__0___redArg___closed__3; @@ -15804,7 +15804,7 @@ if (x_232 == 0) { uint8_t x_233; lean_inc(x_231); -x_233 = l_Lean_Syntax_matchesNull(x_231, x_220); +x_233 = l_Lean_Syntax_matchesNull(x_231, x_218); if (x_233 == 0) { lean_object* x_234; @@ -15813,10 +15813,10 @@ lean_dec(x_224); lean_dec_ref(x_223); lean_dec(x_222); lean_dec(x_221); +lean_dec(x_220); lean_dec(x_219); -lean_dec(x_218); lean_dec(x_217); -lean_dec(x_215); +lean_dec(x_216); x_234 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___redArg(x_225, x_226, x_227, x_228); lean_dec(x_228); lean_dec_ref(x_227); @@ -15834,10 +15834,10 @@ lean_ctor_set(x_236, 0, x_235); x_184 = x_215; x_185 = x_216; x_186 = x_217; -x_187 = x_218; -x_188 = x_221; -x_189 = x_220; -x_190 = x_219; +x_187 = x_221; +x_188 = x_218; +x_189 = x_219; +x_190 = x_220; x_191 = x_236; x_192 = x_222; x_193 = x_223; @@ -15858,10 +15858,10 @@ x_237 = lean_box(0); x_184 = x_215; x_185 = x_216; x_186 = x_217; -x_187 = x_218; -x_188 = x_221; -x_189 = x_220; -x_190 = x_219; +x_187 = x_221; +x_188 = x_218; +x_189 = x_219; +x_190 = x_220; x_191 = x_237; x_192 = x_222; x_193 = x_223; @@ -15960,12 +15960,12 @@ x_263 = l_Lean_Syntax_getArg(x_259, x_141); lean_dec(x_259); x_264 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_264, 0, x_263); -x_215 = x_239; -x_216 = x_253; -x_217 = x_255; -x_218 = x_258; -x_219 = x_250; -x_220 = x_248; +x_215 = x_253; +x_216 = x_239; +x_217 = x_250; +x_218 = x_248; +x_219 = x_258; +x_220 = x_255; x_221 = x_264; x_222 = x_240; x_223 = x_241; @@ -15983,12 +15983,12 @@ else lean_object* x_265; lean_dec(x_259); x_265 = lean_box(0); -x_215 = x_239; -x_216 = x_253; -x_217 = x_255; -x_218 = x_258; -x_219 = x_250; -x_220 = x_248; +x_215 = x_253; +x_216 = x_239; +x_217 = x_250; +x_218 = x_248; +x_219 = x_258; +x_220 = x_255; x_221 = x_265; x_222 = x_240; x_223 = x_241; @@ -16009,17 +16009,17 @@ goto block_238; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; x_88 = l_Array_append___redArg(x_82, x_87); lean_dec_ref(x_87); -lean_inc(x_79); +lean_inc(x_84); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_79); -lean_ctor_set(x_89, 1, x_84); +lean_ctor_set(x_89, 0, x_84); +lean_ctor_set(x_89, 1, x_79); lean_ctor_set(x_89, 2, x_88); x_90 = l_Lean_Elab_Term_CollectPatternVars_collect___lam__0___closed__33; -lean_inc(x_79); +lean_inc(x_84); x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_79); +lean_ctor_set(x_91, 0, x_84); lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_Syntax_node6(x_79, x_78, x_83, x_80, x_86, x_85, x_89, x_91); +x_92 = l_Lean_Syntax_node6(x_84, x_78, x_80, x_83, x_81, x_86, x_89, x_91); x_93 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_93, 0, x_92); return x_93; @@ -16027,52 +16027,52 @@ return x_93; block_113: { lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_inc_ref(x_98); -x_105 = l_Array_append___redArg(x_98, x_104); +lean_inc_ref(x_99); +x_105 = l_Array_append___redArg(x_99, x_104); lean_dec_ref(x_104); -lean_inc(x_102); lean_inc(x_95); +lean_inc(x_101); x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_95); -lean_ctor_set(x_106, 1, x_102); +lean_ctor_set(x_106, 0, x_101); +lean_ctor_set(x_106, 1, x_95); lean_ctor_set(x_106, 2, x_105); -lean_inc(x_95); -x_107 = l_Lean_Syntax_node1(x_95, x_100, x_106); -if (lean_obj_tag(x_99) == 0) +lean_inc(x_101); +x_107 = l_Lean_Syntax_node1(x_101, x_103, x_106); +if (lean_obj_tag(x_96) == 0) { lean_object* x_108; x_108 = l_Lean_Elab_Term_CollectPatternVars_instInhabitedState_default___closed__1; x_79 = x_95; -x_80 = x_96; -x_81 = lean_box(0); -x_82 = x_98; -x_83 = x_101; -x_84 = x_102; -x_85 = x_107; -x_86 = x_103; +x_80 = x_97; +x_81 = x_98; +x_82 = x_99; +x_83 = x_100; +x_84 = x_101; +x_85 = lean_box(0); +x_86 = x_107; x_87 = x_108; goto block_94; } else { lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = lean_ctor_get(x_99, 0); +x_109 = lean_ctor_get(x_96, 0); lean_inc(x_109); -lean_dec_ref(x_99); +lean_dec_ref(x_96); x_110 = l_Lean_Elab_Term_CollectPatternVars_collect___lam__0___closed__34; -lean_inc(x_95); +lean_inc(x_101); x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_95); +lean_ctor_set(x_111, 0, x_101); lean_ctor_set(x_111, 1, x_110); x_112 = l_Array_mkArray2___redArg(x_111, x_109); x_79 = x_95; -x_80 = x_96; -x_81 = lean_box(0); -x_82 = x_98; -x_83 = x_101; -x_84 = x_102; -x_85 = x_107; -x_86 = x_103; +x_80 = x_97; +x_81 = x_98; +x_82 = x_99; +x_83 = x_100; +x_84 = x_101; +x_85 = lean_box(0); +x_86 = x_107; x_87 = x_112; goto block_94; } @@ -16080,41 +16080,41 @@ goto block_94; block_138: { 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_inc_ref(x_117); -x_126 = l_Array_append___redArg(x_117, x_125); +lean_inc_ref(x_119); +x_126 = l_Array_append___redArg(x_119, x_125); lean_dec_ref(x_125); -lean_inc(x_123); lean_inc(x_114); +lean_inc(x_122); x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_114); -lean_ctor_set(x_127, 1, x_123); +lean_ctor_set(x_127, 0, x_122); +lean_ctor_set(x_127, 1, x_114); lean_ctor_set(x_127, 2, x_126); -x_128 = l_Lean_Syntax_SepArray_ofElems(x_124, x_115); -lean_dec_ref(x_115); -lean_inc_ref(x_117); -x_129 = l_Array_append___redArg(x_117, x_128); +x_128 = l_Lean_Syntax_SepArray_ofElems(x_118, x_121); +lean_dec_ref(x_121); +lean_inc_ref(x_119); +x_129 = l_Array_append___redArg(x_119, x_128); lean_dec_ref(x_128); -lean_inc(x_123); lean_inc(x_114); +lean_inc(x_122); x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_114); -lean_ctor_set(x_130, 1, x_123); +lean_ctor_set(x_130, 0, x_122); +lean_ctor_set(x_130, 1, x_114); lean_ctor_set(x_130, 2, x_129); -lean_inc(x_114); -x_131 = l_Lean_Syntax_node1(x_114, x_122, x_130); +lean_inc(x_122); +x_131 = l_Lean_Syntax_node1(x_122, x_117, x_130); if (lean_obj_tag(x_120) == 0) { lean_object* x_132; x_132 = l_Lean_Elab_Term_CollectPatternVars_instInhabitedState_default___closed__1; x_95 = x_114; -x_96 = x_127; -x_97 = lean_box(0); -x_98 = x_117; -x_99 = x_118; -x_100 = x_119; -x_101 = x_121; -x_102 = x_123; -x_103 = x_131; +x_96 = x_115; +x_97 = x_116; +x_98 = x_131; +x_99 = x_119; +x_100 = x_127; +x_101 = x_122; +x_102 = lean_box(0); +x_103 = x_124; x_104 = x_132; goto block_113; } @@ -16132,14 +16132,14 @@ lean_ctor_set(x_136, 0, x_134); lean_ctor_set(x_136, 1, x_135); x_137 = l_Array_mkArray1___redArg(x_136); x_95 = x_114; -x_96 = x_127; -x_97 = lean_box(0); -x_98 = x_117; -x_99 = x_118; -x_100 = x_119; -x_101 = x_121; -x_102 = x_123; -x_103 = x_131; +x_96 = x_115; +x_97 = x_116; +x_98 = x_131; +x_99 = x_119; +x_100 = x_127; +x_101 = x_122; +x_102 = lean_box(0); +x_103 = x_124; x_104 = x_137; goto block_113; } diff --git a/stage0/stdlib/Lean/Elab/Quotation.c b/stage0/stdlib/Lean/Elab/Quotation.c index 7f69a2037117..d48ffc492d5d 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -8041,7 +8041,7 @@ return x_47; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_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; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; 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; 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_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t x_187; size_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; uint8_t x_255; uint8_t x_256; uint8_t x_288; uint8_t x_289; uint8_t x_292; +lean_object* x_48; lean_object* x_49; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_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; uint8_t x_130; lean_object* x_131; 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_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; size_t 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_object* x_197; uint8_t x_255; uint8_t x_256; uint8_t x_288; uint8_t x_289; uint8_t x_292; x_48 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___lam__0___boxed), 8, 0); x_49 = lean_array_uget(x_7, x_9); x_57 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__8; @@ -8111,15 +8111,15 @@ return x_55; block_119: { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_59); -x_71 = l_Lean_Elab_Term_Quotation_mkTuple___redArg(x_60, x_59); +lean_inc_ref(x_66); +x_71 = l_Lean_Elab_Term_Quotation_mkTuple___redArg(x_59, x_66); x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); lean_dec_ref(x_71); -x_73 = lean_ctor_get(x_59, 5); -x_74 = lean_ctor_get(x_59, 10); -x_75 = lean_ctor_get(x_59, 11); -x_76 = l_Lean_SourceInfo_fromRef(x_73, x_64); +x_73 = lean_ctor_get(x_66, 5); +x_74 = lean_ctor_get(x_66, 10); +x_75 = lean_ctor_get(x_66, 11); +x_76 = l_Lean_SourceInfo_fromRef(x_73, x_68); x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; x_78 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__1; x_79 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__2; @@ -8189,8 +8189,8 @@ lean_ctor_set(x_107, 0, x_76); lean_ctor_set(x_107, 1, x_106); x_108 = lean_box(0); x_109 = lean_unsigned_to_nat(0u); -x_110 = lean_array_get(x_108, x_63, x_109); -lean_dec_ref(x_63); +x_110 = lean_array_get(x_108, x_65, x_109); +lean_dec_ref(x_65); lean_inc(x_76); x_111 = l_Lean_Syntax_node4(x_76, x_102, x_103, x_105, x_107, x_110); lean_inc(x_76); @@ -8205,7 +8205,7 @@ x_115 = l_Lean_Syntax_node3(x_76, x_86, x_98, x_112, x_114); lean_inc(x_76); x_116 = l_Lean_Syntax_node2(x_76, x_57, x_115, x_69); x_117 = l_Lean_Syntax_node2(x_76, x_77, x_85, x_116); -x_118 = lean_apply_9(x_66, x_62, x_117, x_65, x_58, x_68, x_67, x_59, x_61, lean_box(0)); +x_118 = lean_apply_9(x_63, x_58, x_117, x_61, x_67, x_64, x_62, x_66, x_60, lean_box(0)); x_24 = x_118; goto block_36; } @@ -8215,17 +8215,17 @@ lean_object* x_132; x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); lean_dec_ref(x_131); -x_58 = x_120; -x_59 = x_121; +x_58 = x_121; +x_59 = x_120; x_60 = x_122; x_61 = x_123; -x_62 = x_125; -x_63 = x_124; +x_62 = x_124; +x_63 = x_125; x_64 = x_126; x_65 = x_128; x_66 = x_127; -x_67 = x_130; -x_68 = x_129; +x_67 = x_129; +x_68 = x_130; x_69 = x_132; x_70 = lean_box(0); goto block_119; @@ -8233,8 +8233,8 @@ goto block_119; block_163: { lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -lean_inc_ref(x_136); -x_149 = l_Array_toSubarray___redArg(x_136, x_147, x_148); +lean_inc_ref(x_134); +x_149 = l_Array_toSubarray___redArg(x_134, x_147, x_148); x_150 = lean_ctor_get(x_149, 0); lean_inc_ref(x_150); x_151 = lean_ctor_get(x_149, 1); @@ -8254,18 +8254,18 @@ if (x_155 == 0) lean_dec(x_153); lean_dec(x_151); lean_dec_ref(x_150); -x_58 = x_142; -x_59 = x_135; -x_60 = x_136; -x_61 = x_137; -x_62 = x_138; +x_58 = x_135; +x_59 = x_134; +x_60 = x_143; +x_61 = x_136; +x_62 = x_144; x_63 = x_139; x_64 = x_140; -x_65 = x_144; +x_65 = x_145; x_66 = x_141; -x_67 = x_145; -x_68 = x_146; -x_69 = x_134; +x_67 = x_146; +x_68 = x_142; +x_69 = x_137; x_70 = lean_box(0); goto block_119; } @@ -8276,21 +8276,21 @@ x_156 = lean_usize_of_nat(x_153); lean_dec(x_153); x_157 = lean_usize_of_nat(x_151); lean_dec(x_151); -lean_inc_ref(x_135); +lean_inc_ref(x_141); lean_inc_ref(x_3); -x_158 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_140, x_3, x_150, x_156, x_157, x_134, x_135); +x_158 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_142, x_3, x_150, x_156, x_157, x_137, x_141); lean_dec_ref(x_150); -x_120 = x_142; +x_120 = x_134; x_121 = x_135; -x_122 = x_136; -x_123 = x_137; -x_124 = x_139; -x_125 = x_138; +x_122 = x_143; +x_123 = x_136; +x_124 = x_144; +x_125 = x_139; x_126 = x_140; x_127 = x_141; -x_128 = x_144; +x_128 = x_145; x_129 = x_146; -x_130 = x_145; +x_130 = x_142; x_131 = x_158; goto block_133; } @@ -8305,18 +8305,18 @@ if (x_159 == 0) lean_dec(x_152); lean_dec(x_151); lean_dec_ref(x_150); -x_58 = x_142; -x_59 = x_135; -x_60 = x_136; -x_61 = x_137; -x_62 = x_138; +x_58 = x_135; +x_59 = x_134; +x_60 = x_143; +x_61 = x_136; +x_62 = x_144; x_63 = x_139; x_64 = x_140; -x_65 = x_144; +x_65 = x_145; x_66 = x_141; -x_67 = x_145; -x_68 = x_146; -x_69 = x_134; +x_67 = x_146; +x_68 = x_142; +x_69 = x_137; x_70 = lean_box(0); goto block_119; } @@ -8327,21 +8327,21 @@ x_160 = lean_usize_of_nat(x_152); lean_dec(x_152); x_161 = lean_usize_of_nat(x_151); lean_dec(x_151); -lean_inc_ref(x_135); +lean_inc_ref(x_141); lean_inc_ref(x_3); -x_162 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_140, x_3, x_150, x_160, x_161, x_134, x_135); +x_162 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_142, x_3, x_150, x_160, x_161, x_137, x_141); lean_dec_ref(x_150); -x_120 = x_142; +x_120 = x_134; x_121 = x_135; -x_122 = x_136; -x_123 = x_137; -x_124 = x_139; -x_125 = x_138; +x_122 = x_143; +x_123 = x_136; +x_124 = x_144; +x_125 = x_139; x_126 = x_140; x_127 = x_141; -x_128 = x_144; +x_128 = x_145; x_129 = x_146; -x_130 = x_145; +x_130 = x_142; x_131 = x_162; goto block_133; } @@ -8360,19 +8360,19 @@ x_182 = lean_nat_dec_le(x_180, x_178); if (x_182 == 0) { lean_dec(x_180); -x_134 = x_177; -x_135 = x_173; -x_136 = x_164; -x_137 = x_174; -x_138 = x_166; -x_139 = x_165; -x_140 = x_167; -x_141 = x_168; -x_142 = x_170; -x_143 = lean_box(0); -x_144 = x_169; -x_145 = x_172; -x_146 = x_171; +x_134 = x_164; +x_135 = x_165; +x_136 = x_169; +x_137 = x_177; +x_138 = lean_box(0); +x_139 = x_166; +x_140 = x_171; +x_141 = x_173; +x_142 = x_168; +x_143 = x_174; +x_144 = x_172; +x_145 = x_167; +x_146 = x_170; x_147 = x_181; x_148 = x_178; goto block_163; @@ -8380,19 +8380,19 @@ goto block_163; else { lean_dec(x_178); -x_134 = x_177; -x_135 = x_173; -x_136 = x_164; -x_137 = x_174; -x_138 = x_166; -x_139 = x_165; -x_140 = x_167; -x_141 = x_168; -x_142 = x_170; -x_143 = lean_box(0); -x_144 = x_169; -x_145 = x_172; -x_146 = x_171; +x_134 = x_164; +x_135 = x_165; +x_136 = x_169; +x_137 = x_177; +x_138 = lean_box(0); +x_139 = x_166; +x_140 = x_171; +x_141 = x_173; +x_142 = x_168; +x_143 = x_174; +x_144 = x_172; +x_145 = x_167; +x_146 = x_170; x_147 = x_181; x_148 = x_180; goto block_163; @@ -8400,12 +8400,12 @@ goto block_163; } block_254: { -if (lean_obj_tag(x_184) == 0) +if (lean_obj_tag(x_187) == 0) { -x_164 = x_185; -x_165 = x_186; -x_166 = x_190; -x_167 = x_187; +x_164 = x_184; +x_165 = x_190; +x_166 = x_185; +x_167 = x_186; x_168 = x_189; x_169 = x_191; x_170 = x_192; @@ -8419,9 +8419,9 @@ goto block_183; else { lean_object* x_198; -x_198 = lean_ctor_get(x_184, 0); +x_198 = lean_ctor_get(x_187, 0); lean_inc(x_198); -lean_dec_ref(x_184); +lean_dec_ref(x_187); if (lean_obj_tag(x_198) == 1) { lean_object* x_199; @@ -8437,10 +8437,10 @@ x_202 = lean_string_dec_eq(x_200, x_201); lean_dec_ref(x_200); if (x_202 == 0) { -x_164 = x_185; -x_165 = x_186; -x_166 = x_190; -x_167 = x_187; +x_164 = x_184; +x_165 = x_190; +x_166 = x_185; +x_167 = x_186; x_168 = x_189; x_169 = x_191; x_170 = x_192; @@ -8457,7 +8457,7 @@ lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; x_203 = lean_ctor_get(x_195, 5); x_204 = lean_ctor_get(x_195, 10); x_205 = lean_ctor_get(x_195, 11); -x_206 = l_Lean_SourceInfo_fromRef(x_203, x_187); +x_206 = l_Lean_SourceInfo_fromRef(x_203, x_189); x_207 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__0; x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__1; x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__2; @@ -8473,11 +8473,11 @@ x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_206); lean_ctor_set(x_214, 1, x_57); lean_ctor_set(x_214, 2, x_213); -x_215 = lean_array_size(x_185); -lean_inc_ref(x_185); +x_215 = lean_array_size(x_184); +lean_inc_ref(x_184); lean_inc_ref(x_214); lean_inc(x_206); -x_216 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__10(x_207, x_208, x_209, x_206, x_214, x_215, x_188, x_185); +x_216 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__10(x_207, x_208, x_209, x_206, x_214, x_215, x_188, x_184); x_217 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__39; x_218 = l_Lean_mkSepArray(x_216, x_217); lean_dec_ref(x_216); @@ -8500,11 +8500,11 @@ lean_inc(x_206); x_226 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_226, 0, x_206); lean_ctor_set(x_226, 1, x_225); -lean_inc_ref(x_185); +lean_inc_ref(x_184); lean_inc(x_206); lean_inc(x_205); lean_inc(x_204); -x_227 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__11(x_207, x_208, x_209, x_204, x_205, x_206, x_57, x_215, x_188, x_185); +x_227 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__11(x_207, x_208, x_209, x_204, x_205, x_206, x_57, x_215, x_188, x_184); x_228 = l_Lean_mkSepArray(x_227, x_217); lean_dec_ref(x_227); x_229 = l_Array_append___redArg(x_213, x_228); @@ -8527,7 +8527,7 @@ lean_inc_ref(x_226); lean_inc(x_206); x_235 = l_Lean_Syntax_node4(x_206, x_224, x_226, x_231, x_233, x_234); lean_inc(x_206); -x_236 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__12(x_207, x_208, x_209, x_4, x_206, x_215, x_188, x_185); +x_236 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__12(x_207, x_208, x_209, x_4, x_206, x_215, x_188, x_184); x_237 = l_Lean_mkSepArray(x_236, x_217); lean_dec_ref(x_236); x_238 = l_Array_append___redArg(x_213, x_237); @@ -8568,7 +8568,7 @@ lean_inc(x_206); x_251 = l_Lean_Syntax_node1(x_206, x_223, x_250); lean_inc_ref(x_214); x_252 = l_Lean_Syntax_node6(x_206, x_211, x_212, x_214, x_214, x_220, x_222, x_251); -x_253 = lean_apply_9(x_189, x_190, x_252, x_191, x_192, x_193, x_194, x_195, x_196, lean_box(0)); +x_253 = lean_apply_9(x_185, x_190, x_252, x_191, x_192, x_193, x_194, x_195, x_196, lean_box(0)); x_24 = x_253; goto block_36; } @@ -8576,10 +8576,10 @@ goto block_36; else { lean_dec_ref(x_198); -x_164 = x_185; -x_165 = x_186; -x_166 = x_190; -x_167 = x_187; +x_164 = x_184; +x_165 = x_190; +x_166 = x_185; +x_167 = x_186; x_168 = x_189; x_169 = x_191; x_170 = x_192; @@ -8594,10 +8594,10 @@ goto block_183; else { lean_dec(x_198); -x_164 = x_185; -x_165 = x_186; -x_166 = x_190; -x_167 = x_187; +x_164 = x_184; +x_165 = x_190; +x_166 = x_185; +x_167 = x_186; x_168 = x_189; x_169 = x_191; x_170 = x_192; @@ -8686,12 +8686,12 @@ lean_inc(x_14); lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); -x_184 = x_269; -x_185 = x_268; +x_184 = x_268; +x_185 = x_271; x_186 = x_266; -x_187 = x_255; +x_187 = x_269; x_188 = x_264; -x_189 = x_271; +x_189 = x_255; x_190 = x_10; x_191 = x_11; x_192 = x_12; @@ -9463,7 +9463,7 @@ return x_47; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t 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_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; 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_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; uint8_t x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; size_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_object* x_197; uint8_t x_255; uint8_t x_256; uint8_t x_288; uint8_t x_289; uint8_t x_292; +lean_object* x_48; lean_object* x_49; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_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; uint8_t x_130; lean_object* x_131; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_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_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; size_t x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; uint8_t x_255; uint8_t x_256; uint8_t x_288; uint8_t x_289; uint8_t x_292; x_48 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___lam__0___boxed), 8, 0); x_49 = lean_array_uget(x_7, x_9); x_57 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__8; @@ -9533,15 +9533,15 @@ return x_55; block_119: { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; 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_inc_ref(x_68); -x_71 = l_Lean_Elab_Term_Quotation_mkTuple___redArg(x_61, x_68); +lean_inc_ref(x_61); +x_71 = l_Lean_Elab_Term_Quotation_mkTuple___redArg(x_63, x_61); x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); lean_dec_ref(x_71); -x_73 = lean_ctor_get(x_68, 5); -x_74 = lean_ctor_get(x_68, 10); -x_75 = lean_ctor_get(x_68, 11); -x_76 = l_Lean_SourceInfo_fromRef(x_73, x_58); +x_73 = lean_ctor_get(x_61, 5); +x_74 = lean_ctor_get(x_61, 10); +x_75 = lean_ctor_get(x_61, 11); +x_76 = l_Lean_SourceInfo_fromRef(x_73, x_68); x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; x_78 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__1; x_79 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__2; @@ -9627,7 +9627,7 @@ x_115 = l_Lean_Syntax_node3(x_76, x_86, x_98, x_112, x_114); lean_inc(x_76); x_116 = l_Lean_Syntax_node2(x_76, x_57, x_115, x_69); x_117 = l_Lean_Syntax_node2(x_76, x_77, x_85, x_116); -x_118 = lean_apply_9(x_67, x_66, x_117, x_65, x_63, x_62, x_59, x_68, x_60, lean_box(0)); +x_118 = lean_apply_9(x_60, x_59, x_117, x_66, x_65, x_67, x_58, x_61, x_62, lean_box(0)); x_24 = x_118; goto block_36; } @@ -9637,15 +9637,15 @@ lean_object* x_132; x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); lean_dec_ref(x_131); -x_58 = x_120; -x_59 = x_121; -x_60 = x_123; -x_61 = x_122; +x_58 = x_123; +x_59 = x_122; +x_60 = x_121; +x_61 = x_120; x_62 = x_124; -x_63 = x_126; -x_64 = x_125; -x_65 = x_127; -x_66 = x_128; +x_63 = x_125; +x_64 = x_126; +x_65 = x_128; +x_66 = x_127; x_67 = x_129; x_68 = x_130; x_69 = x_132; @@ -9655,8 +9655,8 @@ goto block_119; block_163: { lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -lean_inc_ref(x_134); -x_149 = l_Array_toSubarray___redArg(x_134, x_147, x_148); +lean_inc_ref(x_135); +x_149 = l_Array_toSubarray___redArg(x_135, x_147, x_148); x_150 = lean_ctor_get(x_149, 0); lean_inc_ref(x_150); x_151 = lean_ctor_get(x_149, 1); @@ -9676,17 +9676,17 @@ if (x_155 == 0) lean_dec(x_153); lean_dec(x_151); lean_dec_ref(x_150); -x_58 = x_140; -x_59 = x_141; -x_60 = x_142; -x_61 = x_134; -x_62 = x_135; -x_63 = x_144; -x_64 = x_137; +x_58 = x_139; +x_59 = x_134; +x_60 = x_140; +x_61 = x_141; +x_62 = x_142; +x_63 = x_135; +x_64 = x_144; x_65 = x_145; -x_66 = x_138; +x_66 = x_136; x_67 = x_146; -x_68 = x_139; +x_68 = x_137; x_69 = x_143; x_70 = lean_box(0); goto block_119; @@ -9698,21 +9698,21 @@ x_156 = lean_usize_of_nat(x_153); lean_dec(x_153); x_157 = lean_usize_of_nat(x_151); lean_dec(x_151); -lean_inc_ref(x_139); +lean_inc_ref(x_141); lean_inc_ref(x_3); -x_158 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_140, x_3, x_150, x_156, x_157, x_143, x_139); +x_158 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_137, x_3, x_150, x_156, x_157, x_143, x_141); lean_dec_ref(x_150); -x_120 = x_140; -x_121 = x_141; +x_120 = x_141; +x_121 = x_140; x_122 = x_134; -x_123 = x_142; -x_124 = x_135; -x_125 = x_137; +x_123 = x_139; +x_124 = x_142; +x_125 = x_135; x_126 = x_144; -x_127 = x_145; -x_128 = x_138; +x_127 = x_136; +x_128 = x_145; x_129 = x_146; -x_130 = x_139; +x_130 = x_137; x_131 = x_158; goto block_133; } @@ -9727,17 +9727,17 @@ if (x_159 == 0) lean_dec(x_152); lean_dec(x_151); lean_dec_ref(x_150); -x_58 = x_140; -x_59 = x_141; -x_60 = x_142; -x_61 = x_134; -x_62 = x_135; -x_63 = x_144; -x_64 = x_137; +x_58 = x_139; +x_59 = x_134; +x_60 = x_140; +x_61 = x_141; +x_62 = x_142; +x_63 = x_135; +x_64 = x_144; x_65 = x_145; -x_66 = x_138; +x_66 = x_136; x_67 = x_146; -x_68 = x_139; +x_68 = x_137; x_69 = x_143; x_70 = lean_box(0); goto block_119; @@ -9749,21 +9749,21 @@ x_160 = lean_usize_of_nat(x_152); lean_dec(x_152); x_161 = lean_usize_of_nat(x_151); lean_dec(x_151); -lean_inc_ref(x_139); +lean_inc_ref(x_141); lean_inc_ref(x_3); -x_162 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_140, x_3, x_150, x_160, x_161, x_143, x_139); +x_162 = l___private_Init_Data_Array_Basic_0__Array_foldrMUnsafe_fold___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__9___redArg(x_137, x_3, x_150, x_160, x_161, x_143, x_141); lean_dec_ref(x_150); -x_120 = x_140; -x_121 = x_141; +x_120 = x_141; +x_121 = x_140; x_122 = x_134; -x_123 = x_142; -x_124 = x_135; -x_125 = x_137; +x_123 = x_139; +x_124 = x_142; +x_125 = x_135; x_126 = x_144; -x_127 = x_145; -x_128 = x_138; +x_127 = x_136; +x_128 = x_145; x_129 = x_146; -x_130 = x_139; +x_130 = x_137; x_131 = x_162; goto block_133; } @@ -9773,8 +9773,8 @@ goto block_133; { lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; x_176 = lean_box(0); -x_177 = l_Array_back_x21___redArg(x_176, x_165); -x_178 = lean_array_get_size(x_165); +x_177 = l_Array_back_x21___redArg(x_176, x_166); +x_178 = lean_array_get_size(x_166); x_179 = lean_unsigned_to_nat(1u); x_180 = lean_nat_sub(x_178, x_179); x_181 = lean_unsigned_to_nat(0u); @@ -9783,18 +9783,18 @@ if (x_182 == 0) { lean_dec(x_180); x_134 = x_165; -x_135 = x_171; -x_136 = lean_box(0); -x_137 = x_166; -x_138 = x_167; -x_139 = x_173; +x_135 = x_166; +x_136 = x_169; +x_137 = x_168; +x_138 = lean_box(0); +x_139 = x_172; x_140 = x_164; -x_141 = x_172; +x_141 = x_173; x_142 = x_174; x_143 = x_177; -x_144 = x_170; -x_145 = x_169; -x_146 = x_168; +x_144 = x_167; +x_145 = x_170; +x_146 = x_171; x_147 = x_181; x_148 = x_178; goto block_163; @@ -9803,18 +9803,18 @@ else { lean_dec(x_178); x_134 = x_165; -x_135 = x_171; -x_136 = lean_box(0); -x_137 = x_166; -x_138 = x_167; -x_139 = x_173; +x_135 = x_166; +x_136 = x_169; +x_137 = x_168; +x_138 = lean_box(0); +x_139 = x_172; x_140 = x_164; -x_141 = x_172; +x_141 = x_173; x_142 = x_174; x_143 = x_177; -x_144 = x_170; -x_145 = x_169; -x_146 = x_168; +x_144 = x_167; +x_145 = x_170; +x_146 = x_171; x_147 = x_181; x_148 = x_180; goto block_163; @@ -9822,13 +9822,13 @@ goto block_163; } block_254: { -if (lean_obj_tag(x_186) == 0) +if (lean_obj_tag(x_188) == 0) { -x_164 = x_184; -x_165 = x_185; -x_166 = x_187; -x_167 = x_190; -x_168 = x_188; +x_164 = x_185; +x_165 = x_190; +x_166 = x_186; +x_167 = x_187; +x_168 = x_189; x_169 = x_191; x_170 = x_192; x_171 = x_193; @@ -9841,9 +9841,9 @@ goto block_183; else { lean_object* x_198; -x_198 = lean_ctor_get(x_186, 0); +x_198 = lean_ctor_get(x_188, 0); lean_inc(x_198); -lean_dec_ref(x_186); +lean_dec_ref(x_188); if (lean_obj_tag(x_198) == 1) { lean_object* x_199; @@ -9859,11 +9859,11 @@ x_202 = lean_string_dec_eq(x_200, x_201); lean_dec_ref(x_200); if (x_202 == 0) { -x_164 = x_184; -x_165 = x_185; -x_166 = x_187; -x_167 = x_190; -x_168 = x_188; +x_164 = x_185; +x_165 = x_190; +x_166 = x_186; +x_167 = x_187; +x_168 = x_189; x_169 = x_191; x_170 = x_192; x_171 = x_193; @@ -9879,7 +9879,7 @@ lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; x_203 = lean_ctor_get(x_195, 5); x_204 = lean_ctor_get(x_195, 10); x_205 = lean_ctor_get(x_195, 11); -x_206 = l_Lean_SourceInfo_fromRef(x_203, x_184); +x_206 = l_Lean_SourceInfo_fromRef(x_203, x_189); x_207 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__0; x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__1; x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__2; @@ -9895,11 +9895,11 @@ x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_206); lean_ctor_set(x_214, 1, x_57); lean_ctor_set(x_214, 2, x_213); -x_215 = lean_array_size(x_185); -lean_inc_ref(x_185); +x_215 = lean_array_size(x_186); +lean_inc_ref(x_186); lean_inc_ref(x_214); lean_inc(x_206); -x_216 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__10(x_207, x_208, x_209, x_206, x_214, x_215, x_189, x_185); +x_216 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__10(x_207, x_208, x_209, x_206, x_214, x_215, x_184, x_186); x_217 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__39; x_218 = l_Lean_mkSepArray(x_216, x_217); lean_dec_ref(x_216); @@ -9922,11 +9922,11 @@ lean_inc(x_206); x_226 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_226, 0, x_206); lean_ctor_set(x_226, 1, x_225); -lean_inc_ref(x_185); +lean_inc_ref(x_186); lean_inc(x_206); lean_inc(x_205); lean_inc(x_204); -x_227 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__11(x_207, x_208, x_209, x_204, x_205, x_206, x_57, x_215, x_189, x_185); +x_227 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__11(x_207, x_208, x_209, x_204, x_205, x_206, x_57, x_215, x_184, x_186); x_228 = l_Lean_mkSepArray(x_227, x_217); lean_dec_ref(x_227); x_229 = l_Array_append___redArg(x_213, x_228); @@ -9949,7 +9949,7 @@ lean_inc_ref(x_226); lean_inc(x_206); x_235 = l_Lean_Syntax_node4(x_206, x_224, x_226, x_231, x_233, x_234); lean_inc(x_206); -x_236 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__12(x_207, x_208, x_209, x_4, x_206, x_215, x_189, x_185); +x_236 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__12(x_207, x_208, x_209, x_4, x_206, x_215, x_184, x_186); x_237 = l_Lean_mkSepArray(x_236, x_217); lean_dec_ref(x_236); x_238 = l_Array_append___redArg(x_213, x_237); @@ -9990,7 +9990,7 @@ lean_inc(x_206); x_251 = l_Lean_Syntax_node1(x_206, x_223, x_250); lean_inc_ref(x_214); x_252 = l_Lean_Syntax_node6(x_206, x_211, x_212, x_214, x_214, x_220, x_222, x_251); -x_253 = lean_apply_9(x_188, x_190, x_252, x_191, x_192, x_193, x_194, x_195, x_196, lean_box(0)); +x_253 = lean_apply_9(x_185, x_190, x_252, x_191, x_192, x_193, x_194, x_195, x_196, lean_box(0)); x_24 = x_253; goto block_36; } @@ -9998,11 +9998,11 @@ goto block_36; else { lean_dec_ref(x_198); -x_164 = x_184; -x_165 = x_185; -x_166 = x_187; -x_167 = x_190; -x_168 = x_188; +x_164 = x_185; +x_165 = x_190; +x_166 = x_186; +x_167 = x_187; +x_168 = x_189; x_169 = x_191; x_170 = x_192; x_171 = x_193; @@ -10016,11 +10016,11 @@ goto block_183; else { lean_dec(x_198); -x_164 = x_184; -x_165 = x_185; -x_166 = x_187; -x_167 = x_190; -x_168 = x_188; +x_164 = x_185; +x_165 = x_190; +x_166 = x_186; +x_167 = x_187; +x_168 = x_189; x_169 = x_191; x_170 = x_192; x_171 = x_193; @@ -10108,12 +10108,12 @@ lean_inc(x_14); lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); -x_184 = x_255; -x_185 = x_268; -x_186 = x_269; +x_184 = x_264; +x_185 = x_271; +x_186 = x_268; x_187 = x_266; -x_188 = x_271; -x_189 = x_264; +x_188 = x_269; +x_189 = x_255; x_190 = x_10; x_191 = x_11; x_192 = x_12; @@ -11110,7 +11110,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_checkPrivateInPublic___at___00Lean_resolveGlobalName___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__19_spec__19_spec__20_spec__20_spec__20___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; uint8_t x_53; uint8_t 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; uint8_t x_82; lean_object* 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; 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; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_50; lean_object* x_51; uint8_t 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; lean_object* x_78; uint8_t x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* 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; 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) @@ -11147,13 +11147,13 @@ x_26 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_10); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_16); -lean_ctor_set(x_27, 1, x_12); -lean_ctor_set(x_27, 2, x_14); +lean_ctor_set(x_27, 0, x_12); +lean_ctor_set(x_27, 1, x_16); +lean_ctor_set(x_27, 2, x_13); lean_ctor_set(x_27, 3, x_11); lean_ctor_set(x_27, 4, x_26); -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, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_14); 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); @@ -11192,13 +11192,13 @@ x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_10); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_16); -lean_ctor_set(x_43, 1, x_12); -lean_ctor_set(x_43, 2, x_14); +lean_ctor_set(x_43, 0, x_12); +lean_ctor_set(x_43, 1, x_16); +lean_ctor_set(x_43, 2, x_13); lean_ctor_set(x_43, 3, x_11); lean_ctor_set(x_43, 4, x_42); -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, x_15); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_14); 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); @@ -11228,25 +11228,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_52); -x_62 = l_Lean_FileMap_toPosition(x_52, x_51); -lean_dec(x_51); -x_63 = l_Lean_FileMap_toPosition(x_52, x_57); +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_dec(x_57); x_64 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_64, 0, x_63); x_65 = l_panic___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice_spec__0___closed__0; -if (x_54 == 0) +if (x_52 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); x_10 = x_61; x_11 = x_65; -x_12 = x_62; -x_13 = x_53; -x_14 = x_64; -x_15 = x_55; -x_16 = x_56; +x_12 = x_51; +x_13 = x_64; +x_14 = x_54; +x_15 = x_56; +x_16 = x_62; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -11263,7 +11263,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_56); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -11274,11 +11274,11 @@ else lean_free_object(x_59); x_10 = x_61; x_11 = x_65; -x_12 = x_62; -x_13 = x_53; -x_14 = x_64; -x_15 = x_55; -x_16 = x_56; +x_12 = x_51; +x_13 = x_64; +x_14 = x_54; +x_15 = x_56; +x_16 = x_62; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -11292,24 +11292,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_52); -x_69 = l_Lean_FileMap_toPosition(x_52, x_51); -lean_dec(x_51); -x_70 = l_Lean_FileMap_toPosition(x_52, x_57); +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_dec(x_57); x_71 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_71, 0, x_70); x_72 = l_panic___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice_spec__0___closed__0; -if (x_54 == 0) +if (x_52 == 0) { lean_dec_ref(x_50); x_10 = x_68; x_11 = x_72; -x_12 = x_69; -x_13 = x_53; -x_14 = x_71; -x_15 = x_55; -x_16 = x_56; +x_12 = x_51; +x_13 = x_71; +x_14 = x_54; +x_15 = x_56; +x_16 = x_69; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -11326,7 +11326,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_56); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -11337,11 +11337,11 @@ else { x_10 = x_68; x_11 = x_72; -x_12 = x_69; -x_13 = x_53; -x_14 = x_71; -x_15 = x_55; -x_16 = x_56; +x_12 = x_51; +x_13 = x_71; +x_14 = x_54; +x_15 = x_56; +x_16 = x_69; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -11353,18 +11353,18 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_80, x_79); -lean_dec(x_80); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_83, x_82); +lean_dec(x_83); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_84; -x_52 = x_78; -x_53 = x_79; +x_51 = x_78; +x_52 = x_79; +x_53 = x_84; x_54 = x_81; -x_55 = x_82; -x_56 = x_83; +x_55 = x_80; +x_56 = x_82; x_57 = x_84; goto block_76; } @@ -11375,12 +11375,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_84; -x_52 = x_78; -x_53 = x_79; +x_51 = x_78; +x_52 = x_79; +x_53 = x_84; x_54 = x_81; -x_55 = x_82; -x_56 = x_83; +x_55 = x_80; +x_56 = x_82; x_57 = x_86; goto block_76; } @@ -11388,9 +11388,9 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_92); -lean_dec(x_92); -x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_90); +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; @@ -11398,10 +11398,10 @@ x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; x_78 = x_89; x_79 = x_90; -x_80 = x_95; -x_81 = x_91; -x_82 = x_94; -x_83 = x_93; +x_80 = x_91; +x_81 = x_94; +x_82 = x_92; +x_83 = x_95; x_84 = x_97; goto block_87; } @@ -11414,10 +11414,10 @@ lean_dec_ref(x_96); x_77 = x_88; x_78 = x_89; x_79 = x_90; -x_80 = x_95; -x_81 = x_91; -x_82 = x_94; -x_83 = x_93; +x_80 = x_91; +x_81 = x_94; +x_82 = x_92; +x_83 = x_95; x_84 = x_98; goto block_87; } @@ -11426,22 +11426,22 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_102; -x_89 = x_101; -x_90 = x_106; -x_91 = x_103; -x_92 = x_104; +x_88 = x_101; +x_89 = x_102; +x_90 = x_103; +x_91 = x_104; +x_92 = x_106; x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_102; -x_89 = x_101; -x_90 = x_106; -x_91 = x_103; -x_92 = x_104; +x_88 = x_101; +x_89 = x_102; +x_90 = x_103; +x_91 = x_104; +x_92 = x_106; x_93 = x_105; x_94 = x_100; goto block_99; @@ -11466,14 +11466,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(x_113); lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_117; +lean_inc_ref(x_110); +x_101 = x_117; +x_102 = x_110; x_103 = x_114; -x_104 = x_113; -x_105 = x_110; +x_104 = x_111; +x_105 = x_113; x_106 = x_109; x_107 = x_119; goto block_108; @@ -11483,14 +11483,14 @@ else lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_log___at___00Lean_logWarning___at___00Lean_checkPrivateInPublic___at___00Lean_resolveGlobalName___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__19_spec__19_spec__20_spec__20_spec__20___redArg___closed__0; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1_spec__1_spec__2_spec__2(x_112, x_120); -lean_inc_ref(x_110); lean_inc(x_113); lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_117; +lean_inc_ref(x_110); +x_101 = x_117; +x_102 = x_110; x_103 = x_114; -x_104 = x_113; -x_105 = x_110; +x_104 = x_111; +x_105 = x_113; x_106 = x_109; x_107 = x_121; goto block_108; @@ -12704,7 +12704,7 @@ return x_9; } case 1: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_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; uint8_t x_76; uint8_t x_77; 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; uint8_t x_90; uint8_t x_91; 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; uint8_t x_219; 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_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_273; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_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; uint8_t x_77; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; uint8_t x_91; 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; uint8_t x_219; 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_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_273; x_10 = lean_ctor_get(x_1, 1); x_273 = l_Lean_Syntax_antiquotKind_x3f(x_1); if (lean_obj_tag(x_273) == 0) @@ -12796,26 +12796,26 @@ return x_281; block_33: { lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; -x_24 = l_Lean_Syntax_getArgs(x_19); +x_24 = l_Lean_Syntax_getArgs(x_21); x_25 = lean_array_size(x_24); x_26 = 0; -lean_inc_ref(x_20); -x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15(x_23, x_10, x_16, x_22, x_19, x_21, x_24, x_25, x_26, x_11, x_15, x_17, x_18, x_14, x_20, x_12); +lean_inc_ref(x_22); +x_27 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15(x_23, x_10, x_18, x_20, x_21, x_17, x_24, x_25, x_26, x_12, x_15, x_19, x_11, x_14, x_22, x_13); lean_dec_ref(x_24); -lean_dec(x_19); +lean_dec(x_21); 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_Elab_Term_Quotation_ArrayStxBuilder_mkNode___redArg(x_28, x_10, x_20); +x_29 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___redArg(x_28, x_10, x_22); return x_29; } else { uint8_t x_30; -lean_dec_ref(x_20); +lean_dec_ref(x_22); lean_dec(x_10); x_30 = !lean_is_exclusive(x_27); if (x_30 == 0) @@ -12840,7 +12840,7 @@ if (x_42 == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_inc(x_10); -x_43 = lean_st_ref_get(x_34); +x_43 = lean_st_ref_get(x_35); x_44 = lean_ctor_get(x_43, 0); lean_inc_ref(x_44); lean_dec_ref(x_43); @@ -12848,39 +12848,39 @@ x_45 = l_Lean_Syntax_unescapeAntiquot(x_1); x_46 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty; x_47 = l___private_Init_Meta_Defs_0__Lean_quoteArray_go___at___00__private_Init_Meta_Defs_0__Lean_quoteArray___at___00Lean_Elab_Term_Quotation_ArrayStxBuilder_build_spec__0_spec__1___closed__0; x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1; -x_49 = l_Lean_Environment_contains(x_44, x_48, x_41); +x_49 = l_Lean_Environment_contains(x_44, x_48, x_39); if (x_49 == 0) { lean_object* x_50; x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; -x_11 = x_46; -x_12 = x_34; -x_13 = lean_box(0); -x_14 = x_35; +x_11 = x_34; +x_12 = x_46; +x_13 = x_35; +x_14 = x_37; x_15 = x_36; -x_16 = x_47; -x_17 = x_38; -x_18 = x_39; -x_19 = x_45; -x_20 = x_40; -x_21 = x_42; +x_16 = lean_box(0); +x_17 = x_42; +x_18 = x_47; +x_19 = x_38; +x_20 = x_39; +x_21 = x_45; x_22 = x_41; x_23 = x_50; goto block_33; } else { -x_11 = x_46; -x_12 = x_34; -x_13 = lean_box(0); -x_14 = x_35; +x_11 = x_34; +x_12 = x_46; +x_13 = x_35; +x_14 = x_37; x_15 = x_36; -x_16 = x_47; -x_17 = x_38; -x_18 = x_39; -x_19 = x_45; -x_20 = x_40; -x_21 = x_42; +x_16 = lean_box(0); +x_17 = x_42; +x_18 = x_47; +x_19 = x_38; +x_20 = x_39; +x_21 = x_45; x_22 = x_41; x_23 = x_48; goto block_33; @@ -12890,10 +12890,10 @@ else { lean_object* x_51; lean_object* x_52; x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_52 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_51, x_36, x_38, x_39, x_35, x_40, x_34); -lean_dec(x_34); +x_52 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_51, x_36, x_38, x_34, x_37, x_41, x_35); lean_dec(x_35); -lean_dec_ref(x_39); +lean_dec(x_37); +lean_dec_ref(x_34); lean_dec(x_38); lean_dec_ref(x_1); return x_52; @@ -12905,14 +12905,14 @@ uint8_t x_63; x_63 = l_Lean_Syntax_isAntiquotSplice(x_1); if (x_63 == 0) { -lean_dec_ref(x_59); +lean_dec_ref(x_56); x_34 = x_54; x_35 = x_55; -x_36 = x_56; -x_37 = lean_box(0); -x_38 = x_58; -x_39 = x_60; -x_40 = x_61; +x_36 = x_58; +x_37 = x_57; +x_38 = x_59; +x_39 = x_61; +x_40 = lean_box(0); x_41 = x_62; x_42 = x_63; goto block_53; @@ -12921,15 +12921,15 @@ else { lean_object* x_64; lean_object* x_65; uint8_t x_66; x_64 = lean_box(0); -x_65 = lean_apply_1(x_59, x_64); +x_65 = lean_apply_1(x_56, x_64); x_66 = lean_unbox(x_65); x_34 = x_54; x_35 = x_55; -x_36 = x_56; -x_37 = lean_box(0); -x_38 = x_58; -x_39 = x_60; -x_40 = x_61; +x_36 = x_58; +x_37 = x_57; +x_38 = x_59; +x_39 = x_61; +x_40 = lean_box(0); x_41 = x_62; x_42 = x_66; goto block_53; @@ -12942,24 +12942,24 @@ if (x_77 == 0) x_54 = x_68; x_55 = x_69; x_56 = x_70; -x_57 = lean_box(0); -x_58 = x_72; +x_57 = x_72; +x_58 = x_71; x_59 = x_73; -x_60 = x_74; -x_61 = x_75; +x_60 = lean_box(0); +x_61 = x_74; x_62 = x_76; goto block_67; } else { lean_object* x_78; lean_object* x_79; -lean_dec_ref(x_73); +lean_dec_ref(x_70); x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_79 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_78, x_70, x_72, x_74, x_69, x_75, x_68); -lean_dec(x_68); +x_79 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_78, x_71, x_73, x_68, x_72, x_76, x_69); lean_dec(x_69); -lean_dec_ref(x_74); lean_dec(x_72); +lean_dec_ref(x_68); +lean_dec(x_73); lean_dec_ref(x_1); return x_79; } @@ -12973,13 +12973,13 @@ x_92 = l_Lean_Syntax_isAntiquotSuffixSplice(x_1); if (x_92 == 0) { x_68 = x_81; -x_69 = x_82; -x_70 = x_83; -x_71 = lean_box(0); -x_72 = x_86; +x_69 = x_83; +x_70 = x_84; +x_71 = x_86; +x_72 = x_85; x_73 = x_87; -x_74 = x_88; -x_75 = x_89; +x_74 = x_89; +x_75 = lean_box(0); x_76 = x_90; x_77 = x_92; goto block_80; @@ -12994,13 +12994,13 @@ lean_dec(x_94); if (x_95 == 0) { x_68 = x_81; -x_69 = x_82; -x_70 = x_83; -x_71 = lean_box(0); -x_72 = x_86; +x_69 = x_83; +x_70 = x_84; +x_71 = x_86; +x_72 = x_85; x_73 = x_87; -x_74 = x_88; -x_75 = x_89; +x_74 = x_89; +x_75 = lean_box(0); x_76 = x_90; x_77 = x_92; goto block_80; @@ -13008,12 +13008,12 @@ goto block_80; else { x_54 = x_81; -x_55 = x_82; -x_56 = x_83; -x_57 = lean_box(0); +x_55 = x_83; +x_56 = x_84; +x_57 = x_85; x_58 = x_86; x_59 = x_87; -x_60 = x_88; +x_60 = lean_box(0); x_61 = x_89; x_62 = x_90; goto block_67; @@ -13023,17 +13023,17 @@ goto block_67; else { lean_object* x_96; lean_object* x_97; -lean_dec_ref(x_87); +lean_dec_ref(x_84); x_96 = lean_unsigned_to_nat(0u); x_97 = l_Lean_Syntax_getArg(x_1, x_96); if (lean_obj_tag(x_97) == 2) { uint8_t x_98; -lean_dec_ref(x_88); -lean_dec(x_86); -lean_dec_ref(x_83); -lean_dec(x_82); -lean_dec(x_81); +lean_dec(x_87); +lean_dec_ref(x_86); +lean_dec(x_85); +lean_dec(x_83); +lean_dec_ref(x_81); x_98 = !lean_is_exclusive(x_97); if (x_98 == 0) { @@ -13041,14 +13041,14 @@ lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; l x_99 = lean_ctor_get(x_97, 1); x_100 = lean_ctor_get(x_97, 0); lean_dec(x_100); -x_101 = lean_ctor_get(x_89, 5); +x_101 = lean_ctor_get(x_90, 5); lean_inc(x_101); -x_102 = lean_ctor_get(x_89, 10); +x_102 = lean_ctor_get(x_90, 10); lean_inc(x_102); -x_103 = lean_ctor_get(x_89, 11); +x_103 = lean_ctor_get(x_90, 11); lean_inc(x_103); -lean_dec_ref(x_89); -x_104 = l_Lean_SourceInfo_fromRef(x_101, x_85); +lean_dec_ref(x_90); +x_104 = l_Lean_SourceInfo_fromRef(x_101, x_82); lean_dec(x_101); x_105 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; @@ -13160,14 +13160,14 @@ lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; x_154 = lean_ctor_get(x_97, 1); lean_inc(x_154); lean_dec(x_97); -x_155 = lean_ctor_get(x_89, 5); +x_155 = lean_ctor_get(x_90, 5); lean_inc(x_155); -x_156 = lean_ctor_get(x_89, 10); +x_156 = lean_ctor_get(x_90, 10); lean_inc(x_156); -x_157 = lean_ctor_get(x_89, 11); +x_157 = lean_ctor_get(x_90, 11); lean_inc(x_157); -lean_dec_ref(x_89); -x_158 = l_Lean_SourceInfo_fromRef(x_155, x_85); +lean_dec_ref(x_90); +x_158 = l_Lean_SourceInfo_fromRef(x_155, x_82); lean_dec(x_155); x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; @@ -13280,11 +13280,11 @@ else lean_object* x_209; lean_object* x_210; lean_dec(x_97); x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__36; -x_210 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_209, x_83, x_86, x_88, x_82, x_89, x_81); -lean_dec(x_81); -lean_dec(x_82); -lean_dec_ref(x_88); -lean_dec(x_86); +x_210 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_1, x_209, x_86, x_87, x_81, x_85, x_90, x_83); +lean_dec(x_83); +lean_dec(x_85); +lean_dec_ref(x_81); +lean_dec(x_87); lean_dec_ref(x_1); return x_210; } @@ -13304,15 +13304,15 @@ x_223 = l_Lean_Syntax_isTokenAntiquot(x_1); if (x_223 == 0) { x_81 = x_212; -x_82 = x_213; -x_83 = x_214; -x_84 = lean_box(0); -x_85 = x_219; -x_86 = x_216; -x_87 = x_221; -x_88 = x_217; -x_89 = x_218; -x_90 = x_222; +x_82 = x_219; +x_83 = x_213; +x_84 = x_221; +x_85 = x_214; +x_86 = x_215; +x_87 = x_216; +x_88 = lean_box(0); +x_89 = x_222; +x_90 = x_218; x_91 = x_223; goto block_211; } @@ -13322,15 +13322,15 @@ lean_object* x_224; uint8_t x_225; x_224 = lean_box(0); x_225 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lam__0(x_1, x_219, x_224); x_81 = x_212; -x_82 = x_213; -x_83 = x_214; -x_84 = lean_box(0); -x_85 = x_219; -x_86 = x_216; -x_87 = x_221; -x_88 = x_217; -x_89 = x_218; -x_90 = x_222; +x_82 = x_219; +x_83 = x_213; +x_84 = x_221; +x_85 = x_214; +x_86 = x_215; +x_87 = x_216; +x_88 = lean_box(0); +x_89 = x_222; +x_90 = x_218; x_91 = x_225; goto block_211; } @@ -13341,10 +13341,10 @@ if (x_234 == 0) { x_212 = x_227; x_213 = x_228; -x_214 = x_229; -x_215 = lean_box(0); +x_214 = x_230; +x_215 = x_229; x_216 = x_231; -x_217 = x_232; +x_217 = lean_box(0); x_218 = x_233; x_219 = x_234; goto block_226; @@ -13352,11 +13352,11 @@ goto block_226; else { lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -lean_dec_ref(x_232); lean_dec(x_231); +lean_dec(x_230); lean_dec_ref(x_229); lean_dec(x_228); -lean_dec(x_227); +lean_dec_ref(x_227); x_235 = lean_ctor_get(x_233, 5); lean_inc(x_235); x_236 = lean_ctor_get(x_233, 10); @@ -13410,12 +13410,12 @@ lean_inc_ref(x_1); x_268 = l_Lean_Syntax_isAntiquots(x_1); if (x_268 == 0) { -x_227 = x_266; -x_228 = x_264; +x_227 = x_263; +x_228 = x_266; x_229 = x_261; -x_230 = lean_box(0); +x_230 = x_264; x_231 = x_262; -x_232 = x_263; +x_232 = lean_box(0); x_233 = x_265; x_234 = x_268; goto block_260; @@ -13429,12 +13429,12 @@ x_270 = l_Lean_Syntax_isEscapedAntiquot(x_269); lean_dec(x_269); if (x_270 == 0) { -x_227 = x_266; -x_228 = x_264; +x_227 = x_263; +x_228 = x_266; x_229 = x_261; -x_230 = lean_box(0); +x_230 = x_264; x_231 = x_262; -x_232 = x_263; +x_232 = lean_box(0); x_233 = x_265; x_234 = x_268; goto block_260; @@ -13443,12 +13443,12 @@ else { uint8_t x_271; x_271 = 0; -x_212 = x_266; -x_213 = x_264; -x_214 = x_261; -x_215 = lean_box(0); +x_212 = x_263; +x_213 = x_266; +x_214 = x_264; +x_215 = x_261; x_216 = x_262; -x_217 = x_263; +x_217 = lean_box(0); x_218 = x_265; x_219 = x_271; goto block_226; @@ -24281,8 +24281,8 @@ goto block_321; block_233: { lean_object* x_37; -lean_inc_ref(x_35); -x_37 = lean_apply_7(x_2, x_27, x_25, x_33, x_32, x_35, x_34, lean_box(0)); +lean_inc_ref(x_25); +x_37 = lean_apply_7(x_2, x_35, x_32, x_34, x_29, x_25, x_27, lean_box(0)); if (lean_obj_tag(x_37) == 0) { uint8_t x_38; @@ -24291,11 +24291,11 @@ if (x_38 == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; x_39 = lean_ctor_get(x_37, 0); -x_40 = lean_ctor_get(x_35, 10); +x_40 = lean_ctor_get(x_25, 10); lean_inc(x_40); -x_41 = lean_ctor_get(x_35, 11); +x_41 = lean_ctor_get(x_25, 11); lean_inc(x_41); -lean_dec_ref(x_35); +lean_dec_ref(x_25); x_42 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__36; x_43 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__37; lean_inc(x_39); @@ -24354,7 +24354,7 @@ lean_ctor_set(x_69, 1, x_66); lean_ctor_set(x_69, 2, x_68); lean_ctor_set(x_69, 3, x_19); lean_inc(x_39); -x_70 = l_Lean_Syntax_node3(x_39, x_63, x_30, x_65, x_69); +x_70 = l_Lean_Syntax_node3(x_39, x_63, x_33, x_65, x_69); x_71 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__30; x_72 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__31; lean_inc(x_39); @@ -24410,7 +24410,7 @@ lean_ctor_set(x_95, 1, x_91); lean_ctor_set(x_95, 2, x_93); lean_ctor_set(x_95, 3, x_94); lean_inc(x_39); -x_96 = l_Lean_Syntax_node1(x_39, x_45, x_28); +x_96 = l_Lean_Syntax_node1(x_39, x_45, x_26); lean_inc_ref(x_95); lean_inc(x_39); x_97 = l_Lean_Syntax_node2(x_39, x_62, x_95, x_96); @@ -24473,7 +24473,7 @@ x_122 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_122, 0, x_39); lean_ctor_set(x_122, 1, x_121); lean_inc(x_39); -x_123 = l_Lean_Syntax_node1(x_39, x_45, x_31); +x_123 = l_Lean_Syntax_node1(x_39, x_45, x_30); lean_inc(x_39); x_124 = l_Lean_Syntax_node2(x_39, x_62, x_95, x_123); lean_inc(x_39); @@ -24483,7 +24483,7 @@ x_126 = l_Lean_Syntax_node1(x_39, x_45, x_125); lean_inc_ref(x_87); lean_inc_ref(x_77); lean_inc(x_39); -x_127 = l_Lean_Syntax_node4(x_39, x_75, x_77, x_126, x_87, x_29); +x_127 = l_Lean_Syntax_node4(x_39, x_75, x_77, x_126, x_87, x_28); lean_inc(x_39); x_128 = l_Lean_Syntax_node1(x_39, x_45, x_110); lean_inc(x_39); @@ -24505,11 +24505,11 @@ lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; x_134 = lean_ctor_get(x_37, 0); lean_inc(x_134); lean_dec(x_37); -x_135 = lean_ctor_get(x_35, 10); +x_135 = lean_ctor_get(x_25, 10); lean_inc(x_135); -x_136 = lean_ctor_get(x_35, 11); +x_136 = lean_ctor_get(x_25, 11); lean_inc(x_136); -lean_dec_ref(x_35); +lean_dec_ref(x_25); x_137 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__36; x_138 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__37; lean_inc(x_134); @@ -24568,7 +24568,7 @@ lean_ctor_set(x_164, 1, x_161); lean_ctor_set(x_164, 2, x_163); lean_ctor_set(x_164, 3, x_19); lean_inc(x_134); -x_165 = l_Lean_Syntax_node3(x_134, x_158, x_30, x_160, x_164); +x_165 = l_Lean_Syntax_node3(x_134, x_158, x_33, x_160, x_164); x_166 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__30; x_167 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__31; lean_inc(x_134); @@ -24624,7 +24624,7 @@ lean_ctor_set(x_190, 1, x_186); lean_ctor_set(x_190, 2, x_188); lean_ctor_set(x_190, 3, x_189); lean_inc(x_134); -x_191 = l_Lean_Syntax_node1(x_134, x_140, x_28); +x_191 = l_Lean_Syntax_node1(x_134, x_140, x_26); lean_inc_ref(x_190); lean_inc(x_134); x_192 = l_Lean_Syntax_node2(x_134, x_157, x_190, x_191); @@ -24687,7 +24687,7 @@ x_217 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_217, 0, x_134); lean_ctor_set(x_217, 1, x_216); lean_inc(x_134); -x_218 = l_Lean_Syntax_node1(x_134, x_140, x_31); +x_218 = l_Lean_Syntax_node1(x_134, x_140, x_30); lean_inc(x_134); x_219 = l_Lean_Syntax_node2(x_134, x_157, x_190, x_218); lean_inc(x_134); @@ -24697,7 +24697,7 @@ x_221 = l_Lean_Syntax_node1(x_134, x_140, x_220); lean_inc_ref(x_182); lean_inc_ref(x_172); lean_inc(x_134); -x_222 = l_Lean_Syntax_node4(x_134, x_170, x_172, x_221, x_182, x_29); +x_222 = l_Lean_Syntax_node4(x_134, x_170, x_172, x_221, x_182, x_28); lean_inc(x_134); x_223 = l_Lean_Syntax_node1(x_134, x_140, x_205); lean_inc(x_134); @@ -24719,11 +24719,11 @@ else { uint8_t x_230; lean_dec(x_36); -lean_dec_ref(x_35); -lean_dec(x_31); +lean_dec(x_33); lean_dec(x_30); -lean_dec(x_29); lean_dec(x_28); +lean_dec(x_26); +lean_dec_ref(x_25); lean_dec(x_24); x_230 = !lean_is_exclusive(x_37); if (x_230 == 0) @@ -24757,17 +24757,17 @@ x_251 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_251, 0, x_250); lean_ctor_set(x_251, 1, x_249); lean_ctor_set(x_251, 2, x_235); -x_25 = x_241; -x_26 = lean_box(0); -x_27 = x_240; -x_28 = x_236; -x_29 = x_238; -x_30 = x_237; -x_31 = x_239; -x_32 = x_243; -x_33 = x_242; -x_34 = x_245; -x_35 = x_244; +x_25 = x_244; +x_26 = x_236; +x_27 = x_245; +x_28 = x_238; +x_29 = x_243; +x_30 = x_239; +x_31 = lean_box(0); +x_32 = x_241; +x_33 = x_237; +x_34 = x_242; +x_35 = x_240; x_36 = x_251; goto block_233; } @@ -24777,17 +24777,17 @@ lean_object* x_252; lean_object* x_253; x_252 = lean_box(0); x_253 = lean_array_get(x_252, x_235, x_15); lean_dec_ref(x_235); -x_25 = x_241; -x_26 = lean_box(0); -x_27 = x_240; -x_28 = x_236; -x_29 = x_238; -x_30 = x_237; -x_31 = x_239; -x_32 = x_243; -x_33 = x_242; -x_34 = x_245; -x_35 = x_244; +x_25 = x_244; +x_26 = x_236; +x_27 = x_245; +x_28 = x_238; +x_29 = x_243; +x_30 = x_239; +x_31 = lean_box(0); +x_32 = x_241; +x_33 = x_237; +x_34 = x_242; +x_35 = x_240; x_36 = x_253; goto block_233; } @@ -25379,8 +25379,8 @@ goto block_646; block_558: { lean_object* x_457; -lean_inc_ref(x_455); -x_457 = lean_apply_7(x_2, x_447, x_445, x_453, x_452, x_455, x_454, lean_box(0)); +lean_inc_ref(x_445); +x_457 = lean_apply_7(x_2, x_455, x_452, x_454, x_449, x_445, x_447, lean_box(0)); if (lean_obj_tag(x_457) == 0) { lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; 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_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; @@ -25393,11 +25393,11 @@ if (lean_is_exclusive(x_457)) { lean_dec_ref(x_457); x_459 = lean_box(0); } -x_460 = lean_ctor_get(x_455, 10); +x_460 = lean_ctor_get(x_445, 10); lean_inc(x_460); -x_461 = lean_ctor_get(x_455, 11); +x_461 = lean_ctor_get(x_445, 11); lean_inc(x_461); -lean_dec_ref(x_455); +lean_dec_ref(x_445); x_462 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__36; x_463 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__37; lean_inc(x_458); @@ -25456,7 +25456,7 @@ lean_ctor_set(x_489, 1, x_486); lean_ctor_set(x_489, 2, x_488); lean_ctor_set(x_489, 3, x_19); lean_inc(x_458); -x_490 = l_Lean_Syntax_node3(x_458, x_483, x_450, x_485, x_489); +x_490 = l_Lean_Syntax_node3(x_458, x_483, x_453, x_485, x_489); x_491 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__30; x_492 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___closed__31; lean_inc(x_458); @@ -25512,7 +25512,7 @@ lean_ctor_set(x_515, 1, x_511); lean_ctor_set(x_515, 2, x_513); lean_ctor_set(x_515, 3, x_514); lean_inc(x_458); -x_516 = l_Lean_Syntax_node1(x_458, x_465, x_448); +x_516 = l_Lean_Syntax_node1(x_458, x_465, x_446); lean_inc_ref(x_515); lean_inc(x_458); x_517 = l_Lean_Syntax_node2(x_458, x_482, x_515, x_516); @@ -25575,7 +25575,7 @@ x_542 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_542, 0, x_458); lean_ctor_set(x_542, 1, x_541); lean_inc(x_458); -x_543 = l_Lean_Syntax_node1(x_458, x_465, x_451); +x_543 = l_Lean_Syntax_node1(x_458, x_465, x_450); lean_inc(x_458); x_544 = l_Lean_Syntax_node2(x_458, x_482, x_515, x_543); lean_inc(x_458); @@ -25585,7 +25585,7 @@ x_546 = l_Lean_Syntax_node1(x_458, x_465, x_545); lean_inc_ref(x_507); lean_inc_ref(x_497); lean_inc(x_458); -x_547 = l_Lean_Syntax_node4(x_458, x_495, x_497, x_546, x_507, x_449); +x_547 = l_Lean_Syntax_node4(x_458, x_495, x_497, x_546, x_507, x_448); lean_inc(x_458); x_548 = l_Lean_Syntax_node1(x_458, x_465, x_530); lean_inc(x_458); @@ -25610,11 +25610,11 @@ else { lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_dec(x_456); -lean_dec_ref(x_455); -lean_dec(x_451); +lean_dec(x_453); lean_dec(x_450); -lean_dec(x_449); lean_dec(x_448); +lean_dec(x_446); +lean_dec_ref(x_445); lean_dec(x_444); x_555 = lean_ctor_get(x_457, 0); lean_inc(x_555); @@ -25649,17 +25649,17 @@ x_576 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_576, 0, x_575); lean_ctor_set(x_576, 1, x_574); lean_ctor_set(x_576, 2, x_560); -x_445 = x_566; -x_446 = lean_box(0); -x_447 = x_565; -x_448 = x_561; -x_449 = x_563; -x_450 = x_562; -x_451 = x_564; -x_452 = x_568; -x_453 = x_567; -x_454 = x_570; -x_455 = x_569; +x_445 = x_569; +x_446 = x_561; +x_447 = x_570; +x_448 = x_563; +x_449 = x_568; +x_450 = x_564; +x_451 = lean_box(0); +x_452 = x_566; +x_453 = x_562; +x_454 = x_567; +x_455 = x_565; x_456 = x_576; goto block_558; } @@ -25669,17 +25669,17 @@ lean_object* x_577; lean_object* x_578; x_577 = lean_box(0); x_578 = lean_array_get(x_577, x_560, x_15); lean_dec_ref(x_560); -x_445 = x_566; -x_446 = lean_box(0); -x_447 = x_565; -x_448 = x_561; -x_449 = x_563; -x_450 = x_562; -x_451 = x_564; -x_452 = x_568; -x_453 = x_567; -x_454 = x_570; -x_455 = x_569; +x_445 = x_569; +x_446 = x_561; +x_447 = x_570; +x_448 = x_563; +x_449 = x_568; +x_450 = x_564; +x_451 = lean_box(0); +x_452 = x_566; +x_453 = x_562; +x_454 = x_567; +x_455 = x_565; x_456 = x_578; goto block_558; } @@ -27872,13 +27872,13 @@ return x_736; { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 0, x_19); lean_ctor_set(x_27, 1, x_26); x_28 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___lam__1___closed__3; x_29 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_3, x_29, x_22, x_19, x_21, x_20, x_24, x_23); +x_30 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_3, x_29, x_22, x_24, x_23, x_25, x_20, x_18); return x_30; } block_48: @@ -27889,14 +27889,14 @@ if (lean_obj_tag(x_32) == 0) { lean_object* x_41; x_41 = 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__15_spec__15___lam__1___closed__6; -x_18 = lean_box(0); -x_19 = x_34; -x_20 = x_36; -x_21 = x_35; +x_18 = x_38; +x_19 = x_40; +x_20 = x_37; +x_21 = lean_box(0); x_22 = x_33; -x_23 = x_38; -x_24 = x_37; -x_25 = x_40; +x_23 = x_35; +x_24 = x_34; +x_25 = x_36; x_26 = x_41; goto block_31; } @@ -27915,14 +27915,14 @@ x_46 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00 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_18 = lean_box(0); -x_19 = x_34; -x_20 = x_36; -x_21 = x_35; +x_18 = x_38; +x_19 = x_40; +x_20 = x_37; +x_21 = lean_box(0); x_22 = x_33; -x_23 = x_38; -x_24 = x_37; -x_25 = x_40; +x_23 = x_35; +x_24 = x_34; +x_25 = x_36; x_26 = x_47; goto block_31; } @@ -28840,7 +28840,7 @@ 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; uint8_t x_94; 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; uint8_t x_115; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; uint8_t 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; uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_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_292; +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; uint8_t x_94; 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; uint8_t x_115; uint8_t x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; uint8_t x_167; uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_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_292; lean_free_object(x_1); lean_dec(x_18); lean_dec(x_17); @@ -28959,34 +28959,34 @@ return x_300; lean_object* x_95; lean_object* x_96; lean_object* x_97; x_95 = lean_box(x_22); x_96 = lean_box(x_94); -lean_inc(x_91); -lean_inc_ref(x_90); -lean_inc(x_89); +lean_inc(x_88); +lean_inc_ref(x_89); +lean_inc(x_90); x_97 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__13___boxed), 9, 8); -lean_closure_set(x_97, 0, x_88); +lean_closure_set(x_97, 0, x_91); lean_closure_set(x_97, 1, x_95); -lean_closure_set(x_97, 2, x_89); +lean_closure_set(x_97, 2, x_90); lean_closure_set(x_97, 3, x_87); -lean_closure_set(x_97, 4, x_90); +lean_closure_set(x_97, 4, x_89); lean_closure_set(x_97, 5, x_96); -lean_closure_set(x_97, 6, x_91); +lean_closure_set(x_97, 6, x_88); lean_closure_set(x_97, 7, x_84); if (x_94 == 0) { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_91); +lean_dec(x_88); x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_89); +lean_ctor_set(x_98, 0, x_90); lean_ctor_set(x_98, 1, x_87); -x_99 = lean_array_get_size(x_90); -lean_dec_ref(x_90); +x_99 = lean_array_get_size(x_89); +lean_dec_ref(x_89); x_100 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_100, 0, x_99); x_101 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_101, 0, x_98); lean_ctor_set(x_101, 1, x_100); -x_9 = x_97; -x_10 = x_92; +x_9 = x_92; +x_10 = x_97; x_11 = lean_box(0); x_12 = x_101; goto block_15; @@ -28994,12 +28994,12 @@ goto block_15; else { lean_object* x_102; -lean_dec_ref(x_90); -lean_dec(x_89); +lean_dec(x_90); +lean_dec_ref(x_89); x_102 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_102, 0, x_91); -x_9 = x_97; -x_10 = x_92; +lean_ctor_set(x_102, 0, x_88); +x_9 = x_92; +x_10 = x_97; x_11 = lean_box(0); x_12 = x_102; goto block_15; @@ -29009,14 +29009,14 @@ goto block_15; { lean_object* x_116; lean_object* x_117; x_116 = lean_unsigned_to_nat(0u); -x_117 = l_Array_findIdx_x3f_loop___redArg(x_108, x_112, x_116); -lean_dec_ref(x_112); +x_117 = l_Array_findIdx_x3f_loop___redArg(x_108, x_114, x_116); +lean_dec_ref(x_114); if (lean_obj_tag(x_117) == 0) { lean_object* x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; size_t x_122; size_t x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -lean_dec(x_113); -lean_dec(x_111); -lean_dec_ref(x_110); +lean_dec(x_112); +lean_dec_ref(x_111); +lean_dec(x_110); x_118 = l_Lean_Syntax_unescapeAntiquot(x_109); lean_inc(x_118); x_119 = l_Lean_Syntax_getKind(x_118); @@ -29042,10 +29042,10 @@ lean_closure_set(x_128, 4, x_118); x_129 = l_Lean_Syntax_isIdent(x_118); if (x_129 == 0) { -x_88 = x_125; -x_89 = x_119; -x_90 = x_124; -x_91 = x_118; +x_88 = x_118; +x_89 = x_124; +x_90 = x_119; +x_91 = x_125; x_92 = x_128; x_93 = lean_box(0); x_94 = x_120; @@ -29053,10 +29053,10 @@ goto block_103; } else { -x_88 = x_125; -x_89 = x_119; -x_90 = x_124; -x_91 = x_118; +x_88 = x_118; +x_89 = x_124; +x_90 = x_119; +x_91 = x_125; x_92 = x_128; x_93 = lean_box(0); x_94 = x_22; @@ -29074,8 +29074,8 @@ lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; x_131 = lean_ctor_get(x_117, 0); x_132 = l_Lean_Syntax_getNumArgs(x_109); lean_dec(x_109); -x_133 = lean_nat_sub(x_132, x_113); -lean_dec(x_113); +x_133 = lean_nat_sub(x_132, x_112); +lean_dec(x_112); lean_dec(x_132); x_134 = lean_nat_sub(x_133, x_131); x_135 = lean_box(x_115); @@ -29083,10 +29083,10 @@ x_136 = lean_box(x_22); lean_inc(x_134); lean_inc(x_131); x_137 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__18___boxed), 8, 7); -lean_closure_set(x_137, 0, x_111); +lean_closure_set(x_137, 0, x_110); lean_closure_set(x_137, 1, x_135); lean_closure_set(x_137, 2, x_131); -lean_closure_set(x_137, 3, x_110); +lean_closure_set(x_137, 3, x_111); lean_closure_set(x_137, 4, x_116); lean_closure_set(x_137, 5, x_136); lean_closure_set(x_137, 6, x_134); @@ -29118,8 +29118,8 @@ lean_inc(x_142); lean_dec(x_117); x_143 = l_Lean_Syntax_getNumArgs(x_109); lean_dec(x_109); -x_144 = lean_nat_sub(x_143, x_113); -lean_dec(x_113); +x_144 = lean_nat_sub(x_143, x_112); +lean_dec(x_112); lean_dec(x_143); x_145 = lean_nat_sub(x_144, x_142); x_146 = lean_box(x_115); @@ -29127,10 +29127,10 @@ x_147 = lean_box(x_22); lean_inc(x_145); lean_inc(x_142); x_148 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__18___boxed), 8, 7); -lean_closure_set(x_148, 0, x_111); +lean_closure_set(x_148, 0, x_110); lean_closure_set(x_148, 1, x_146); lean_closure_set(x_148, 2, x_142); -lean_closure_set(x_148, 3, x_110); +lean_closure_set(x_148, 3, x_111); lean_closure_set(x_148, 4, x_116); lean_closure_set(x_148, 5, x_147); lean_closure_set(x_148, 6, x_145); @@ -29160,24 +29160,24 @@ return x_153; { if (x_167 == 0) { -lean_dec_ref(x_165); +lean_dec(x_165); lean_dec_ref(x_164); -lean_dec(x_163); lean_dec(x_162); lean_dec_ref(x_161); -lean_dec(x_159); +lean_dec(x_160); +lean_dec_ref(x_159); lean_dec_ref(x_19); -if (x_160 == 0) +if (x_166 == 0) { lean_dec_ref(x_106); lean_dec(x_21); -lean_inc_ref(x_155); -x_110 = x_155; -x_111 = x_157; -x_112 = x_155; -x_113 = x_156; -x_114 = lean_box(0); -x_115 = x_160; +lean_inc_ref(x_158); +x_110 = x_156; +x_111 = x_158; +x_112 = x_157; +x_113 = lean_box(0); +x_114 = x_158; +x_115 = x_166; goto block_154; } else @@ -29191,20 +29191,20 @@ if (x_170 == 0) { lean_dec_ref(x_106); lean_dec(x_21); -lean_inc_ref(x_155); -x_110 = x_155; -x_111 = x_157; -x_112 = x_155; -x_113 = x_156; -x_114 = lean_box(0); +lean_inc_ref(x_158); +x_110 = x_156; +x_111 = x_158; +x_112 = x_157; +x_113 = lean_box(0); +x_114 = x_158; x_115 = x_170; goto block_154; } 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; -lean_dec(x_157); -lean_dec_ref(x_155); +lean_dec_ref(x_158); +lean_dec(x_156); lean_dec_ref(x_108); lean_dec_ref(x_84); x_171 = lean_box(x_167); @@ -29214,7 +29214,7 @@ x_173 = lean_box(x_167); x_174 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__21___boxed), 14, 5); lean_closure_set(x_174, 0, x_109); lean_closure_set(x_174, 1, x_172); -lean_closure_set(x_174, 2, x_156); +lean_closure_set(x_174, 2, x_157); lean_closure_set(x_174, 3, x_173); lean_closure_set(x_174, 4, x_104); x_175 = lean_alloc_ctor(3, 1, 0); @@ -29232,9 +29232,9 @@ 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; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec_ref(x_158); lean_dec(x_157); lean_dec(x_156); -lean_dec_ref(x_155); lean_dec_ref(x_108); lean_dec_ref(x_106); lean_dec_ref(x_84); @@ -29255,7 +29255,7 @@ x_188 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAda lean_inc(x_184); x_189 = l_Lean_Syntax_isOfKind(x_184, x_188); x_190 = lean_box(x_189); -x_191 = lean_box(x_158); +x_191 = lean_box(x_155); x_192 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__22___boxed), 17, 9); lean_closure_set(x_192, 0, x_190); lean_closure_set(x_192, 1, x_184); @@ -29266,13 +29266,13 @@ lean_closure_set(x_192, 5, x_185); lean_closure_set(x_192, 6, x_186); lean_closure_set(x_192, 7, x_187); lean_closure_set(x_192, 8, x_182); -x_193 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_192, x_165, x_163, x_161, x_162, x_164, x_159); -lean_dec(x_159); -lean_dec_ref(x_164); +x_193 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_192, x_161, x_160, x_164, x_162, x_159, x_165); +lean_dec(x_165); +lean_dec_ref(x_159); lean_dec(x_162); +lean_dec_ref(x_164); +lean_dec(x_160); lean_dec_ref(x_161); -lean_dec(x_163); -lean_dec_ref(x_165); return x_193; } } @@ -29296,18 +29296,18 @@ x_208 = lean_unsigned_to_nat(1u); x_209 = lean_nat_dec_eq(x_207, x_208); if (x_209 == 0) { -x_155 = x_206; -x_156 = x_208; -x_157 = x_207; -x_158 = x_205; -x_159 = x_196; -x_160 = x_209; -x_161 = x_197; -x_162 = x_199; -x_163 = x_198; +x_155 = x_205; +x_156 = x_207; +x_157 = x_208; +x_158 = x_206; +x_159 = x_197; +x_160 = x_196; +x_161 = x_198; +x_162 = x_200; +x_163 = lean_box(0); x_164 = x_201; -x_165 = x_200; -x_166 = lean_box(0); +x_165 = x_202; +x_166 = x_209; x_167 = x_209; goto block_194; } @@ -29318,18 +29318,18 @@ x_210 = lean_unsigned_to_nat(0u); x_211 = l_Lean_Syntax_getArg(x_109, x_210); x_212 = l_Lean_Syntax_isAntiquotSuffixSplice(x_211); lean_dec(x_211); -x_155 = x_206; -x_156 = x_208; -x_157 = x_207; -x_158 = x_205; -x_159 = x_196; -x_160 = x_209; -x_161 = x_197; -x_162 = x_199; -x_163 = x_198; +x_155 = x_205; +x_156 = x_207; +x_157 = x_208; +x_158 = x_206; +x_159 = x_197; +x_160 = x_196; +x_161 = x_198; +x_162 = x_200; +x_163 = lean_box(0); x_164 = x_201; -x_165 = x_200; -x_166 = lean_box(0); +x_165 = x_202; +x_166 = x_209; x_167 = x_212; goto block_194; } @@ -29343,11 +29343,11 @@ lean_dec_ref(x_84); lean_dec(x_21); lean_dec_ref(x_19); x_213 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_214 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_109, x_213, x_200, x_198, x_197, x_199, x_201, x_196); +x_214 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_109, x_213, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); +lean_dec(x_200); +lean_dec_ref(x_201); lean_dec(x_196); -lean_dec(x_199); -lean_dec_ref(x_197); -lean_dec(x_198); lean_dec(x_109); return x_214; } @@ -29361,11 +29361,11 @@ lean_dec_ref(x_84); lean_dec(x_21); lean_dec_ref(x_19); x_215 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_216 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_109, x_215, x_200, x_198, x_197, x_199, x_201, x_196); +x_216 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_109, x_215, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); +lean_dec(x_200); +lean_dec_ref(x_201); lean_dec(x_196); -lean_dec(x_199); -lean_dec_ref(x_197); -lean_dec(x_198); lean_dec(x_109); return x_216; } @@ -29417,10 +29417,10 @@ x_235 = l_List_all___redArg(x_221, x_86); if (x_235 == 0) { 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_202); lean_dec_ref(x_201); -lean_dec_ref(x_200); -lean_dec(x_199); -lean_dec(x_198); +lean_dec(x_200); +lean_dec_ref(x_198); lean_dec_ref(x_197); lean_dec(x_196); lean_dec_ref(x_19); @@ -29455,13 +29455,13 @@ else lean_object* x_243; lean_free_object(x_218); lean_dec(x_220); -x_243 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_231, x_200, x_198, x_197, x_199, x_201, x_196); -lean_dec(x_196); -lean_dec_ref(x_201); -lean_dec(x_199); +x_243 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_231, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); lean_dec_ref(x_197); -lean_dec(x_198); -lean_dec_ref(x_200); +lean_dec(x_200); +lean_dec_ref(x_201); +lean_dec(x_196); +lean_dec_ref(x_198); return x_243; } } @@ -29472,13 +29472,13 @@ lean_free_object(x_218); lean_dec(x_221); lean_dec(x_220); lean_dec_ref(x_86); -x_244 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_231, x_200, x_198, x_197, x_199, x_201, x_196); -lean_dec(x_196); -lean_dec_ref(x_201); -lean_dec(x_199); +x_244 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_231, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); lean_dec_ref(x_197); -lean_dec(x_198); -lean_dec_ref(x_200); +lean_dec(x_200); +lean_dec_ref(x_201); +lean_dec(x_196); +lean_dec_ref(x_198); return x_244; } } @@ -29521,10 +29521,10 @@ x_260 = l_List_all___redArg(x_246, x_86); if (x_260 == 0) { 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_dec(x_202); lean_dec_ref(x_201); -lean_dec_ref(x_200); -lean_dec(x_199); -lean_dec(x_198); +lean_dec(x_200); +lean_dec_ref(x_198); lean_dec_ref(x_197); lean_dec(x_196); lean_dec_ref(x_19); @@ -29559,13 +29559,13 @@ else { lean_object* x_269; lean_dec(x_245); -x_269 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_256, x_200, x_198, x_197, x_199, x_201, x_196); -lean_dec(x_196); -lean_dec_ref(x_201); -lean_dec(x_199); +x_269 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_256, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); lean_dec_ref(x_197); -lean_dec(x_198); -lean_dec_ref(x_200); +lean_dec(x_200); +lean_dec_ref(x_201); +lean_dec(x_196); +lean_dec_ref(x_198); return x_269; } } @@ -29575,13 +29575,13 @@ lean_object* x_270; lean_dec(x_246); lean_dec(x_245); lean_dec_ref(x_86); -x_270 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_256, x_200, x_198, x_197, x_199, x_201, x_196); -lean_dec(x_196); -lean_dec_ref(x_201); -lean_dec(x_199); +x_270 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_19, x_256, x_198, x_196, x_201, x_200, x_197, x_202); +lean_dec(x_202); lean_dec_ref(x_197); -lean_dec(x_198); -lean_dec_ref(x_200); +lean_dec(x_200); +lean_dec_ref(x_201); +lean_dec(x_196); +lean_dec_ref(x_198); return x_270; } } @@ -29610,13 +29610,13 @@ x_283 = l_Lean_Syntax_isAntiquots(x_109); if (x_283 == 0) { x_195 = x_282; -x_196 = x_277; -x_197 = x_274; -x_198 = x_273; -x_199 = x_275; -x_200 = x_272; -x_201 = x_276; -x_202 = lean_box(0); +x_196 = x_273; +x_197 = x_276; +x_198 = x_272; +x_199 = lean_box(0); +x_200 = x_275; +x_201 = x_274; +x_202 = x_277; x_203 = x_283; goto block_271; } @@ -29630,26 +29630,26 @@ lean_dec(x_284); if (x_285 == 0) { x_195 = x_282; -x_196 = x_277; -x_197 = x_274; -x_198 = x_273; -x_199 = x_275; -x_200 = x_272; -x_201 = x_276; -x_202 = lean_box(0); +x_196 = x_273; +x_197 = x_276; +x_198 = x_272; +x_199 = lean_box(0); +x_200 = x_275; +x_201 = x_274; +x_202 = x_277; x_203 = x_283; goto block_271; } else { x_195 = x_282; -x_196 = x_277; -x_197 = x_274; -x_198 = x_273; -x_199 = x_275; -x_200 = x_272; -x_201 = x_276; -x_202 = lean_box(0); +x_196 = x_273; +x_197 = x_276; +x_198 = x_272; +x_199 = lean_box(0); +x_200 = x_275; +x_201 = x_274; +x_202 = x_277; x_203 = x_282; goto block_271; } @@ -29941,7 +29941,7 @@ return x_355; } else { -lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; 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; uint8_t x_387; lean_object* x_416; lean_object* x_417; lean_object* x_418; uint8_t x_419; lean_object* 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_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; uint8_t x_464; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_528; +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; uint8_t x_366; 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; uint8_t x_387; uint8_t x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; uint8_t x_427; uint8_t x_428; uint8_t 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; uint8_t x_464; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_528; lean_dec(x_302); lean_dec(x_301); x_356 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__9___boxed), 8, 0); @@ -30060,34 +30060,34 @@ return x_536; lean_object* x_367; lean_object* x_368; lean_object* x_369; x_367 = lean_box(x_306); x_368 = lean_box(x_366); -lean_inc(x_363); -lean_inc_ref(x_362); -lean_inc(x_361); +lean_inc(x_360); +lean_inc_ref(x_361); +lean_inc(x_362); x_369 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__13___boxed), 9, 8); -lean_closure_set(x_369, 0, x_360); +lean_closure_set(x_369, 0, x_363); lean_closure_set(x_369, 1, x_367); -lean_closure_set(x_369, 2, x_361); +lean_closure_set(x_369, 2, x_362); lean_closure_set(x_369, 3, x_359); -lean_closure_set(x_369, 4, x_362); +lean_closure_set(x_369, 4, x_361); lean_closure_set(x_369, 5, x_368); -lean_closure_set(x_369, 6, x_363); +lean_closure_set(x_369, 6, x_360); lean_closure_set(x_369, 7, x_356); if (x_366 == 0) { lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; -lean_dec(x_363); +lean_dec(x_360); x_370 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_370, 0, x_361); +lean_ctor_set(x_370, 0, x_362); lean_ctor_set(x_370, 1, x_359); -x_371 = lean_array_get_size(x_362); -lean_dec_ref(x_362); +x_371 = lean_array_get_size(x_361); +lean_dec_ref(x_361); x_372 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_372, 0, x_371); x_373 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_373, 0, x_370); lean_ctor_set(x_373, 1, x_372); -x_9 = x_369; -x_10 = x_364; +x_9 = x_364; +x_10 = x_369; x_11 = lean_box(0); x_12 = x_373; goto block_15; @@ -30095,12 +30095,12 @@ goto block_15; else { lean_object* x_374; -lean_dec_ref(x_362); -lean_dec(x_361); +lean_dec(x_362); +lean_dec_ref(x_361); x_374 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_374, 0, x_363); -x_9 = x_369; -x_10 = x_364; +lean_ctor_set(x_374, 0, x_360); +x_9 = x_364; +x_10 = x_369; x_11 = lean_box(0); x_12 = x_374; goto block_15; @@ -30110,14 +30110,14 @@ goto block_15; { lean_object* x_388; lean_object* x_389; x_388 = lean_unsigned_to_nat(0u); -x_389 = l_Array_findIdx_x3f_loop___redArg(x_380, x_384, x_388); -lean_dec_ref(x_384); +x_389 = l_Array_findIdx_x3f_loop___redArg(x_380, x_386, x_388); +lean_dec_ref(x_386); if (lean_obj_tag(x_389) == 0) { lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; size_t x_394; size_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; uint8_t x_401; -lean_dec(x_385); -lean_dec(x_383); -lean_dec_ref(x_382); +lean_dec(x_384); +lean_dec_ref(x_383); +lean_dec(x_382); x_390 = l_Lean_Syntax_unescapeAntiquot(x_381); lean_inc(x_390); x_391 = l_Lean_Syntax_getKind(x_390); @@ -30143,10 +30143,10 @@ lean_closure_set(x_400, 4, x_390); x_401 = l_Lean_Syntax_isIdent(x_390); if (x_401 == 0) { -x_360 = x_397; -x_361 = x_391; -x_362 = x_396; -x_363 = x_390; +x_360 = x_390; +x_361 = x_396; +x_362 = x_391; +x_363 = x_397; x_364 = x_400; x_365 = lean_box(0); x_366 = x_392; @@ -30154,10 +30154,10 @@ goto block_375; } else { -x_360 = x_397; -x_361 = x_391; -x_362 = x_396; -x_363 = x_390; +x_360 = x_390; +x_361 = x_396; +x_362 = x_391; +x_363 = x_397; x_364 = x_400; x_365 = lean_box(0); x_366 = x_306; @@ -30179,8 +30179,8 @@ if (lean_is_exclusive(x_389)) { } x_404 = l_Lean_Syntax_getNumArgs(x_381); lean_dec(x_381); -x_405 = lean_nat_sub(x_404, x_385); -lean_dec(x_385); +x_405 = lean_nat_sub(x_404, x_384); +lean_dec(x_384); lean_dec(x_404); x_406 = lean_nat_sub(x_405, x_402); x_407 = lean_box(x_387); @@ -30188,10 +30188,10 @@ x_408 = lean_box(x_306); lean_inc(x_406); lean_inc(x_402); x_409 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__18___boxed), 8, 7); -lean_closure_set(x_409, 0, x_383); +lean_closure_set(x_409, 0, x_382); lean_closure_set(x_409, 1, x_407); lean_closure_set(x_409, 2, x_402); -lean_closure_set(x_409, 3, x_382); +lean_closure_set(x_409, 3, x_383); lean_closure_set(x_409, 4, x_388); lean_closure_set(x_409, 5, x_408); lean_closure_set(x_409, 6, x_406); @@ -30225,24 +30225,24 @@ return x_414; { if (x_428 == 0) { -lean_dec_ref(x_426); +lean_dec(x_426); lean_dec_ref(x_425); -lean_dec(x_424); lean_dec(x_423); lean_dec_ref(x_422); -lean_dec(x_420); +lean_dec(x_421); +lean_dec_ref(x_420); lean_dec_ref(x_303); -if (x_421 == 0) +if (x_427 == 0) { lean_dec_ref(x_378); lean_dec(x_305); -lean_inc_ref(x_416); -x_382 = x_416; -x_383 = x_418; -x_384 = x_416; -x_385 = x_417; -x_386 = lean_box(0); -x_387 = x_421; +lean_inc_ref(x_419); +x_382 = x_417; +x_383 = x_419; +x_384 = x_418; +x_385 = lean_box(0); +x_386 = x_419; +x_387 = x_427; goto block_415; } else @@ -30256,20 +30256,20 @@ if (x_431 == 0) { lean_dec_ref(x_378); lean_dec(x_305); -lean_inc_ref(x_416); -x_382 = x_416; -x_383 = x_418; -x_384 = x_416; -x_385 = x_417; -x_386 = lean_box(0); +lean_inc_ref(x_419); +x_382 = x_417; +x_383 = x_419; +x_384 = x_418; +x_385 = lean_box(0); +x_386 = x_419; x_387 = x_431; goto block_415; } else { 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_dec(x_418); -lean_dec_ref(x_416); +lean_dec_ref(x_419); +lean_dec(x_417); lean_dec_ref(x_380); lean_dec_ref(x_356); x_432 = lean_box(x_428); @@ -30279,7 +30279,7 @@ x_434 = lean_box(x_428); x_435 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__21___boxed), 14, 5); lean_closure_set(x_435, 0, x_381); lean_closure_set(x_435, 1, x_433); -lean_closure_set(x_435, 2, x_417); +lean_closure_set(x_435, 2, x_418); lean_closure_set(x_435, 3, x_434); lean_closure_set(x_435, 4, x_376); x_436 = lean_alloc_ctor(3, 1, 0); @@ -30297,9 +30297,9 @@ return x_438; else { lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; uint8_t x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; +lean_dec_ref(x_419); lean_dec(x_418); lean_dec(x_417); -lean_dec_ref(x_416); lean_dec_ref(x_380); lean_dec_ref(x_378); lean_dec_ref(x_356); @@ -30320,7 +30320,7 @@ x_449 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAda lean_inc(x_445); x_450 = l_Lean_Syntax_isOfKind(x_445, x_449); x_451 = lean_box(x_450); -x_452 = lean_box(x_419); +x_452 = lean_box(x_416); x_453 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__22___boxed), 17, 9); lean_closure_set(x_453, 0, x_451); lean_closure_set(x_453, 1, x_445); @@ -30331,13 +30331,13 @@ lean_closure_set(x_453, 5, x_446); lean_closure_set(x_453, 6, x_447); lean_closure_set(x_453, 7, x_448); lean_closure_set(x_453, 8, x_443); -x_454 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_453, x_426, x_424, x_422, x_423, x_425, x_420); -lean_dec(x_420); -lean_dec_ref(x_425); +x_454 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_453, x_422, x_421, x_425, x_423, x_420, x_426); +lean_dec(x_426); +lean_dec_ref(x_420); lean_dec(x_423); +lean_dec_ref(x_425); +lean_dec(x_421); lean_dec_ref(x_422); -lean_dec(x_424); -lean_dec_ref(x_426); return x_454; } } @@ -30361,18 +30361,18 @@ x_469 = lean_unsigned_to_nat(1u); x_470 = lean_nat_dec_eq(x_468, x_469); if (x_470 == 0) { -x_416 = x_467; -x_417 = x_469; -x_418 = x_468; -x_419 = x_466; -x_420 = x_457; -x_421 = x_470; -x_422 = x_458; -x_423 = x_460; -x_424 = x_459; +x_416 = x_466; +x_417 = x_468; +x_418 = x_469; +x_419 = x_467; +x_420 = x_458; +x_421 = x_457; +x_422 = x_459; +x_423 = x_461; +x_424 = lean_box(0); x_425 = x_462; -x_426 = x_461; -x_427 = lean_box(0); +x_426 = x_463; +x_427 = x_470; x_428 = x_470; goto block_455; } @@ -30383,18 +30383,18 @@ x_471 = lean_unsigned_to_nat(0u); x_472 = l_Lean_Syntax_getArg(x_381, x_471); x_473 = l_Lean_Syntax_isAntiquotSuffixSplice(x_472); lean_dec(x_472); -x_416 = x_467; -x_417 = x_469; -x_418 = x_468; -x_419 = x_466; -x_420 = x_457; -x_421 = x_470; -x_422 = x_458; -x_423 = x_460; -x_424 = x_459; +x_416 = x_466; +x_417 = x_468; +x_418 = x_469; +x_419 = x_467; +x_420 = x_458; +x_421 = x_457; +x_422 = x_459; +x_423 = x_461; +x_424 = lean_box(0); x_425 = x_462; -x_426 = x_461; -x_427 = lean_box(0); +x_426 = x_463; +x_427 = x_470; x_428 = x_473; goto block_455; } @@ -30408,11 +30408,11 @@ lean_dec_ref(x_356); lean_dec(x_305); lean_dec_ref(x_303); x_474 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_475 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_381, x_474, x_461, x_459, x_458, x_460, x_462, x_457); +x_475 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_381, x_474, x_459, x_457, x_462, x_461, x_458, x_463); +lean_dec(x_463); +lean_dec(x_461); +lean_dec_ref(x_462); lean_dec(x_457); -lean_dec(x_460); -lean_dec_ref(x_458); -lean_dec(x_459); lean_dec(x_381); return x_475; } @@ -30426,11 +30426,11 @@ lean_dec_ref(x_356); lean_dec(x_305); lean_dec_ref(x_303); x_476 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_477 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_381, x_476, x_461, x_459, x_458, x_460, x_462, x_457); +x_477 = l_Lean_throwErrorAt___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax_spec__1___redArg(x_381, x_476, x_459, x_457, x_462, x_461, x_458, x_463); +lean_dec(x_463); +lean_dec(x_461); +lean_dec_ref(x_462); lean_dec(x_457); -lean_dec(x_460); -lean_dec_ref(x_458); -lean_dec(x_459); lean_dec(x_381); return x_477; } @@ -30488,10 +30488,10 @@ x_496 = l_List_all___redArg(x_481, x_358); if (x_496 == 0) { lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; +lean_dec(x_463); lean_dec_ref(x_462); -lean_dec_ref(x_461); -lean_dec(x_460); -lean_dec(x_459); +lean_dec(x_461); +lean_dec_ref(x_459); lean_dec_ref(x_458); lean_dec(x_457); lean_dec_ref(x_303); @@ -30532,13 +30532,13 @@ else lean_object* x_505; lean_dec(x_482); lean_dec(x_480); -x_505 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_492, x_461, x_459, x_458, x_460, x_462, x_457); -lean_dec(x_457); -lean_dec_ref(x_462); -lean_dec(x_460); +x_505 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_492, x_459, x_457, x_462, x_461, x_458, x_463); +lean_dec(x_463); lean_dec_ref(x_458); -lean_dec(x_459); -lean_dec_ref(x_461); +lean_dec(x_461); +lean_dec_ref(x_462); +lean_dec(x_457); +lean_dec_ref(x_459); return x_505; } } @@ -30549,13 +30549,13 @@ lean_dec(x_482); lean_dec(x_481); lean_dec(x_480); lean_dec_ref(x_358); -x_506 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_492, x_461, x_459, x_458, x_460, x_462, x_457); -lean_dec(x_457); -lean_dec_ref(x_462); -lean_dec(x_460); +x_506 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lam__3(x_303, x_492, x_459, x_457, x_462, x_461, x_458, x_463); +lean_dec(x_463); lean_dec_ref(x_458); -lean_dec(x_459); -lean_dec_ref(x_461); +lean_dec(x_461); +lean_dec_ref(x_462); +lean_dec(x_457); +lean_dec_ref(x_459); return x_506; } } @@ -30583,13 +30583,13 @@ x_519 = l_Lean_Syntax_isAntiquots(x_381); if (x_519 == 0) { x_456 = x_518; -x_457 = x_513; -x_458 = x_510; -x_459 = x_509; -x_460 = x_511; -x_461 = x_508; -x_462 = x_512; -x_463 = lean_box(0); +x_457 = x_509; +x_458 = x_512; +x_459 = x_508; +x_460 = lean_box(0); +x_461 = x_511; +x_462 = x_510; +x_463 = x_513; x_464 = x_519; goto block_507; } @@ -30603,26 +30603,26 @@ lean_dec(x_520); if (x_521 == 0) { x_456 = x_518; -x_457 = x_513; -x_458 = x_510; -x_459 = x_509; -x_460 = x_511; -x_461 = x_508; -x_462 = x_512; -x_463 = lean_box(0); +x_457 = x_509; +x_458 = x_512; +x_459 = x_508; +x_460 = lean_box(0); +x_461 = x_511; +x_462 = x_510; +x_463 = x_513; x_464 = x_519; goto block_507; } else { x_456 = x_518; -x_457 = x_513; -x_458 = x_510; -x_459 = x_509; -x_460 = x_511; -x_461 = x_508; -x_462 = x_512; -x_463 = lean_box(0); +x_457 = x_509; +x_458 = x_512; +x_459 = x_508; +x_460 = lean_box(0); +x_461 = x_511; +x_462 = x_510; +x_463 = x_513; x_464 = x_518; goto block_507; } @@ -30695,8 +30695,8 @@ return x_526; lean_object* x_13; lean_object* x_14; x_13 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 1, x_10); +lean_ctor_set(x_13, 2, x_9); x_14 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_14, 0, x_13); return x_14; @@ -32700,13 +32700,13 @@ if (lean_is_scalar(x_15)) { x_21 = x_15; lean_ctor_set_tag(x_21, 0); } -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_17); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 0, x_16); lean_ctor_set(x_22, 1, x_21); x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_16); +lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); x_3 = x_14; x_4 = x_23; @@ -32834,11 +32834,11 @@ lean_dec_ref(x_42); lean_dec(x_41); lean_dec(x_38); lean_dec(x_36); -x_16 = x_43; -x_17 = x_74; -x_18 = x_39; -x_19 = lean_box(0); -x_20 = x_44; +x_16 = x_44; +x_17 = lean_box(0); +x_18 = x_43; +x_19 = x_39; +x_20 = x_74; goto block_25; } else @@ -32878,11 +32878,11 @@ lean_dec_ref(x_42); lean_dec(x_41); lean_dec(x_38); lean_dec(x_36); -x_16 = x_43; -x_17 = x_74; -x_18 = x_39; -x_19 = lean_box(0); -x_20 = x_44; +x_16 = x_44; +x_17 = lean_box(0); +x_18 = x_43; +x_19 = x_39; +x_20 = x_74; goto block_25; } } @@ -38456,7 +38456,7 @@ return x_4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_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; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +lean_object* x_9; lean_object* x_10; lean_object* x_11; 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_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__0; x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__1; x_52 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lam__0___closed__2; @@ -38478,7 +38478,7 @@ return x_55; } else { -lean_object* x_56; lean_object* x_57; size_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_88; lean_object* x_89; size_t x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; lean_object* x_99; lean_object* x_100; uint8_t x_101; +lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; uint8_t x_93; lean_object* x_99; lean_object* x_100; uint8_t x_101; x_56 = lean_unsigned_to_nat(0u); x_99 = lean_unsigned_to_nat(1u); x_100 = l_Lean_Syntax_getArg(x_1, x_99); @@ -38646,11 +38646,11 @@ x_130 = lean_nat_dec_lt(x_56, x_129); if (x_130 == 0) { lean_dec(x_129); -x_88 = x_126; +x_88 = x_113; x_89 = x_128; -x_90 = x_110; -x_91 = x_127; -x_92 = x_113; +x_90 = x_126; +x_91 = x_110; +x_92 = x_127; x_93 = x_117; goto block_98; } @@ -38659,11 +38659,11 @@ else if (x_130 == 0) { lean_dec(x_129); -x_88 = x_126; +x_88 = x_113; x_89 = x_128; -x_90 = x_110; -x_91 = x_127; -x_92 = x_113; +x_90 = x_126; +x_91 = x_110; +x_92 = x_127; x_93 = x_117; goto block_98; } @@ -38675,20 +38675,20 @@ lean_dec(x_129); x_132 = l___private_Init_Data_Array_Basic_0__Array_anyMUnsafe_any___at___00Lean_Elab_Term_Quotation_match__syntax_expand_spec__12(x_56, x_50, x_51, x_52, x_103, x_107, x_127, x_110, x_131); if (x_132 == 0) { -x_88 = x_126; +x_88 = x_113; x_89 = x_128; -x_90 = x_110; -x_91 = x_127; -x_92 = x_113; +x_90 = x_126; +x_91 = x_110; +x_92 = x_127; x_93 = x_117; goto block_98; } else { -x_57 = x_126; -x_58 = x_110; -x_59 = x_128; -x_60 = x_113; +x_57 = x_113; +x_58 = x_128; +x_59 = x_110; +x_60 = x_126; x_61 = x_127; x_62 = x_2; x_63 = x_3; @@ -38712,8 +38712,8 @@ goto block_87; lean_object* x_69; lean_inc(x_67); lean_inc_ref(x_66); -x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss(x_57, x_62, x_63, x_64, x_65, x_66, x_67); -lean_dec_ref(x_57); +x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss(x_60, x_62, x_63, x_64, x_65, x_66, x_67); +lean_dec_ref(x_60); if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; @@ -38735,19 +38735,19 @@ x_76 = lean_nat_dec_lt(x_56, x_75); if (x_76 == 0) { lean_dec(x_75); -x_9 = x_74; +x_9 = x_58; x_10 = x_64; -x_11 = x_59; -x_12 = x_65; -x_13 = x_73; -x_14 = x_66; -x_15 = x_72; -x_16 = x_61; -x_17 = x_62; +x_11 = x_65; +x_12 = x_66; +x_13 = x_62; +x_14 = x_73; +x_15 = x_57; +x_16 = x_59; +x_17 = x_61; x_18 = x_67; -x_19 = x_58; -x_20 = x_63; -x_21 = x_60; +x_19 = x_63; +x_20 = x_72; +x_21 = x_74; x_22 = lean_box(0); goto block_49; } @@ -38758,19 +38758,19 @@ x_77 = lean_nat_dec_le(x_75, x_75); if (x_77 == 0) { lean_dec(x_75); -x_9 = x_74; +x_9 = x_58; x_10 = x_64; -x_11 = x_59; -x_12 = x_65; -x_13 = x_73; -x_14 = x_66; -x_15 = x_72; -x_16 = x_61; -x_17 = x_62; +x_11 = x_65; +x_12 = x_66; +x_13 = x_62; +x_14 = x_73; +x_15 = x_57; +x_16 = x_59; +x_17 = x_61; x_18 = x_67; -x_19 = x_58; -x_20 = x_63; -x_21 = x_60; +x_19 = x_63; +x_20 = x_72; +x_21 = x_74; x_22 = lean_box(0); goto block_49; } @@ -38786,23 +38786,23 @@ lean_inc(x_65); lean_inc_ref(x_64); lean_inc(x_63); lean_inc_ref(x_62); -x_80 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Term_Quotation_match__syntax_expand_spec__11(x_56, x_61, x_58, x_79, x_78, x_62, x_63, x_64, x_65, x_66, x_67); +x_80 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Elab_Term_Quotation_match__syntax_expand_spec__11(x_56, x_61, x_59, x_79, x_78, x_62, x_63, x_64, x_65, x_66, x_67); if (lean_obj_tag(x_80) == 0) { lean_dec_ref(x_80); -x_9 = x_74; +x_9 = x_58; x_10 = x_64; -x_11 = x_59; -x_12 = x_65; -x_13 = x_73; -x_14 = x_66; -x_15 = x_72; -x_16 = x_61; -x_17 = x_62; +x_11 = x_65; +x_12 = x_66; +x_13 = x_62; +x_14 = x_73; +x_15 = x_57; +x_16 = x_59; +x_17 = x_61; x_18 = x_67; -x_19 = x_58; -x_20 = x_63; -x_21 = x_60; +x_19 = x_63; +x_20 = x_72; +x_21 = x_74; x_22 = lean_box(0); goto block_49; } @@ -38819,8 +38819,8 @@ lean_dec_ref(x_64); lean_dec(x_63); lean_dec_ref(x_62); lean_dec_ref(x_61); -lean_dec_ref(x_60); -lean_dec_ref(x_59); +lean_dec_ref(x_58); +lean_dec_ref(x_57); x_81 = !lean_is_exclusive(x_80); if (x_81 == 0) { @@ -38850,8 +38850,8 @@ lean_dec_ref(x_64); lean_dec(x_63); lean_dec_ref(x_62); lean_dec_ref(x_61); -lean_dec_ref(x_60); -lean_dec_ref(x_59); +lean_dec_ref(x_58); +lean_dec_ref(x_57); x_84 = !lean_is_exclusive(x_69); if (x_84 == 0) { @@ -38874,10 +38874,10 @@ return x_86; if (x_93 == 0) { x_57 = x_88; -x_58 = x_90; -x_59 = x_89; -x_60 = x_92; -x_61 = x_91; +x_58 = x_89; +x_59 = x_91; +x_60 = x_90; +x_61 = x_92; x_62 = x_2; x_63 = x_3; x_64 = x_4; @@ -38891,7 +38891,7 @@ else { lean_object* x_94; uint8_t x_95; lean_dec_ref(x_92); -lean_dec_ref(x_91); +lean_dec_ref(x_90); lean_dec_ref(x_89); lean_dec_ref(x_88); lean_dec(x_7); @@ -38922,38 +38922,38 @@ return x_97; block_49: { lean_object* x_23; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_23 = lean_array_to_list(x_21); -x_24 = lean_array_size(x_16); -x_25 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_Quotation_match__syntax_expand_spec__10(x_24, x_19, x_16); -x_26 = l_Array_zip___redArg(x_25, x_9); -lean_dec_ref(x_9); +x_23 = lean_array_to_list(x_15); +x_24 = lean_array_size(x_17); +x_25 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Lean_Elab_Term_Quotation_match__syntax_expand_spec__10(x_24, x_16, x_17); +x_26 = l_Array_zip___redArg(x_25, x_21); +lean_dec_ref(x_21); lean_dec_ref(x_25); x_27 = lean_array_to_list(x_26); lean_inc(x_18); -lean_inc_ref(x_14); -lean_inc(x_12); +lean_inc_ref(x_12); +lean_inc(x_11); lean_inc_ref(x_10); -lean_inc(x_20); -lean_inc_ref(x_17); -x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch(x_23, x_27, x_17, x_20, x_10, x_12, x_14, x_18); +lean_inc(x_19); +lean_inc_ref(x_13); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch(x_23, x_27, x_13, x_19, x_10, x_11, x_12, x_18); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); lean_dec_ref(x_28); -lean_inc_ref(x_14); -x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts(x_29, x_11, x_15, x_13, x_17, x_20, x_10, x_12, x_14, x_18); -lean_dec(x_20); -lean_dec_ref(x_17); +lean_inc_ref(x_12); +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts(x_29, x_9, x_20, x_14, x_13, x_19, x_10, x_11, x_12, x_18); +lean_dec(x_19); lean_dec_ref(x_13); -lean_dec(x_15); -lean_dec_ref(x_11); +lean_dec_ref(x_14); +lean_dec(x_20); +lean_dec_ref(x_9); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); lean_dec_ref(x_30); x_32 = l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; -x_33 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__2___redArg(x_32, x_14); +x_33 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__2___redArg(x_32, x_12); x_34 = !lean_is_exclusive(x_33); if (x_34 == 0) { @@ -38964,8 +38964,8 @@ lean_dec(x_35); if (x_36 == 0) { lean_dec(x_18); -lean_dec_ref(x_14); -lean_dec(x_12); +lean_dec_ref(x_12); +lean_dec(x_11); lean_dec_ref(x_10); lean_ctor_set(x_33, 0, x_31); return x_33; @@ -38976,10 +38976,10 @@ lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_free_object(x_33); lean_inc(x_31); x_37 = l_Lean_MessageData_ofSyntax(x_31); -x_38 = l_Lean_addTrace___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__3___redArg(x_32, x_37, x_10, x_12, x_14, x_18); +x_38 = l_Lean_addTrace___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__3___redArg(x_32, x_37, x_10, x_11, x_12, x_18); lean_dec(x_18); -lean_dec_ref(x_14); -lean_dec(x_12); +lean_dec_ref(x_12); +lean_dec(x_11); lean_dec_ref(x_10); x_39 = !lean_is_exclusive(x_38); if (x_39 == 0) @@ -39012,8 +39012,8 @@ if (x_43 == 0) { lean_object* x_44; lean_dec(x_18); -lean_dec_ref(x_14); -lean_dec(x_12); +lean_dec_ref(x_12); +lean_dec(x_11); lean_dec_ref(x_10); x_44 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_44, 0, x_31); @@ -39024,10 +39024,10 @@ else lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_inc(x_31); x_45 = l_Lean_MessageData_ofSyntax(x_31); -x_46 = l_Lean_addTrace___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__3___redArg(x_32, x_45, x_10, x_12, x_14, x_18); +x_46 = l_Lean_addTrace___at___00__private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch_spec__3___redArg(x_32, x_45, x_10, x_11, x_12, x_18); lean_dec(x_18); -lean_dec_ref(x_14); -lean_dec(x_12); +lean_dec_ref(x_12); +lean_dec(x_11); lean_dec_ref(x_10); if (lean_is_exclusive(x_46)) { lean_ctor_release(x_46, 0); @@ -39049,14 +39049,14 @@ return x_48; else { lean_dec(x_20); +lean_dec(x_19); lean_dec(x_18); -lean_dec_ref(x_17); -lean_dec(x_15); lean_dec_ref(x_14); lean_dec_ref(x_13); -lean_dec(x_12); -lean_dec_ref(x_11); +lean_dec_ref(x_12); +lean_dec(x_11); lean_dec_ref(x_10); +lean_dec_ref(x_9); return x_28; } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Config.c b/stage0/stdlib/Lean/Elab/Tactic/Config.c index 9f9b17cbaeab..02696fe1a07b 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Config.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Config.c @@ -2987,7 +2987,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0_spec__0_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) { _start: { -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; 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; uint8_t x_80; lean_object* x_81; uint8_t x_82; lean_object* 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; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_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; 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; 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) @@ -3022,15 +3022,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_16); +lean_ctor_set(x_26, 1, x_11); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_13); -lean_ctor_set(x_27, 1, x_14); -lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_11); +lean_ctor_set(x_27, 0, x_12); +lean_ctor_set(x_27, 1, x_15); +lean_ctor_set(x_27, 2, x_14); +lean_ctor_set(x_27, 3, x_10); 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_10); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_13); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_16); 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); @@ -3067,15 +3067,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_16); +lean_ctor_set(x_42, 1, x_11); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_13); -lean_ctor_set(x_43, 1, x_14); -lean_ctor_set(x_43, 2, x_15); -lean_ctor_set(x_43, 3, x_11); +lean_ctor_set(x_43, 0, x_12); +lean_ctor_set(x_43, 1, x_15); +lean_ctor_set(x_43, 2, x_14); +lean_ctor_set(x_43, 3, x_10); 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_10); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_13); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_16); 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); @@ -3105,25 +3105,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_51); -lean_dec(x_51); -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_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0_spec__0_spec__0___redArg___closed__0; -if (x_56 == 0) +if (x_53 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_52; -x_11 = x_65; -x_12 = x_54; -x_13 = x_53; -x_14 = x_62; -x_15 = x_64; -x_16 = x_61; +x_10 = x_65; +x_11 = x_61; +x_12 = x_51; +x_13 = x_52; +x_14 = x_64; +x_15 = x_62; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -3140,7 +3140,7 @@ lean_object* x_67; lean_dec_ref(x_64); lean_dec_ref(x_62); lean_dec(x_61); -lean_dec_ref(x_53); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_67 = lean_box(0); lean_ctor_set(x_59, 0, x_67); @@ -3149,13 +3149,13 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_52; -x_11 = x_65; -x_12 = x_54; -x_13 = x_53; -x_14 = x_62; -x_15 = x_64; -x_16 = x_61; +x_10 = x_65; +x_11 = x_61; +x_12 = x_51; +x_13 = x_52; +x_14 = x_64; +x_15 = x_62; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -3169,24 +3169,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_51); -lean_dec(x_51); -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_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0_spec__0_spec__0___redArg___closed__0; -if (x_56 == 0) +if (x_53 == 0) { lean_dec_ref(x_50); -x_10 = x_52; -x_11 = x_72; -x_12 = x_54; -x_13 = x_53; -x_14 = x_69; -x_15 = x_71; -x_16 = x_68; +x_10 = x_72; +x_11 = x_68; +x_12 = x_51; +x_13 = x_52; +x_14 = x_71; +x_15 = x_69; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -3203,7 +3203,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_53); +lean_dec_ref(x_51); lean_dec_ref(x_7); x_74 = lean_box(0); x_75 = lean_alloc_ctor(0, 1, 0); @@ -3212,13 +3212,13 @@ return x_75; } else { -x_10 = x_52; -x_11 = x_72; -x_12 = x_54; -x_13 = x_53; -x_14 = x_69; -x_15 = x_71; -x_16 = x_68; +x_10 = x_72; +x_11 = x_68; +x_12 = x_51; +x_13 = x_52; +x_14 = x_71; +x_15 = x_69; +x_16 = x_56; x_17 = x_7; x_18 = x_8; x_19 = lean_box(0); @@ -3230,18 +3230,18 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_83, x_80); -lean_dec(x_83); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_82, x_79); +lean_dec(x_82); if (lean_obj_tag(x_85) == 0) { lean_inc(x_84); x_50 = x_77; -x_51 = x_84; -x_52 = x_78; +x_51 = x_78; +x_52 = x_79; x_53 = x_81; x_54 = x_80; -x_55 = x_79; -x_56 = x_82; +x_55 = x_84; +x_56 = x_83; x_57 = x_84; goto block_76; } @@ -3252,12 +3252,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_84; -x_52 = x_78; +x_51 = x_78; +x_52 = x_79; x_53 = x_81; x_54 = x_80; -x_55 = x_79; -x_56 = x_82; +x_55 = x_84; +x_56 = x_83; x_57 = x_86; goto block_76; } @@ -3265,20 +3265,20 @@ goto block_76; block_99: { lean_object* x_95; lean_object* x_96; -x_95 = l_Lean_replaceRef(x_1, x_92); -lean_dec(x_92); +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_94; -x_79 = x_91; -x_80 = x_90; -x_81 = x_89; -x_82 = x_93; -x_83 = x_95; +x_78 = x_89; +x_79 = x_90; +x_80 = x_92; +x_81 = x_91; +x_82 = x_95; +x_83 = x_94; x_84 = x_97; goto block_87; } @@ -3289,12 +3289,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_94; -x_79 = x_91; -x_80 = x_90; -x_81 = x_89; -x_82 = x_93; -x_83 = x_95; +x_78 = x_89; +x_79 = x_90; +x_80 = x_92; +x_81 = x_91; +x_82 = x_95; +x_83 = x_94; x_84 = x_98; goto block_87; } @@ -3303,22 +3303,22 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_103; -x_89 = x_102; +x_88 = x_104; +x_89 = x_101; x_90 = x_106; -x_91 = x_101; -x_92 = x_104; +x_91 = x_103; +x_92 = x_102; x_93 = x_105; x_94 = x_3; goto block_99; } else { -x_88 = x_103; -x_89 = x_102; +x_88 = x_104; +x_89 = x_101; x_90 = x_106; -x_91 = x_101; -x_92 = x_104; +x_91 = x_103; +x_92 = x_102; x_93 = x_105; x_94 = x_100; goto block_99; @@ -3344,13 +3344,13 @@ 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_ref(x_111); -x_101 = x_111; -x_102 = x_110; -x_103 = x_117; -x_104 = x_113; -x_105 = x_114; +lean_inc_ref(x_110); +x_101 = x_110; +x_102 = x_111; +x_103 = x_114; +x_104 = x_117; +x_105 = x_113; x_106 = x_109; x_107 = x_119; goto block_108; @@ -3361,13 +3361,13 @@ lean_object* x_120; uint8_t x_121; x_120 = l_Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0_spec__0_spec__0___redArg___closed__1; x_121 = l_Lean_Option_get___at___00Lean_logAt___at___00Lean_logErrorAt___at___00Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0_spec__0_spec__0_spec__0(x_112, x_120); lean_inc(x_113); -lean_inc_ref(x_110); lean_inc_ref(x_111); -x_101 = x_111; -x_102 = x_110; -x_103 = x_117; -x_104 = x_113; -x_105 = x_114; +lean_inc_ref(x_110); +x_101 = x_110; +x_102 = x_111; +x_103 = x_114; +x_104 = x_117; +x_105 = x_113; x_106 = x_109; x_107 = x_121; goto block_108; @@ -6340,14 +6340,14 @@ lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec(x_1); x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 0, x_31); return x_34; } else { lean_object* x_35; lean_inc_ref(x_14); -x_35 = l_Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0(x_32, x_10, x_11, x_12, x_13, x_14, x_15); +x_35 = l_Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0(x_31, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; lean_object* x_37; @@ -6425,7 +6425,7 @@ lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec(x_1); x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 0, x_31); return x_41; } } @@ -6437,15 +6437,15 @@ if (x_45 == 0) { uint8_t x_46; x_46 = l_Lean_Exception_isRuntime(x_43); -x_31 = lean_box(0); -x_32 = x_43; +x_31 = x_43; +x_32 = lean_box(0); x_33 = x_46; goto block_42; } else { -x_31 = lean_box(0); -x_32 = x_43; +x_31 = x_43; +x_32 = lean_box(0); x_33 = x_45; goto block_42; } @@ -6942,14 +6942,14 @@ lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec(x_1); x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 0, x_31); return x_34; } else { lean_object* x_35; lean_inc_ref(x_14); -x_35 = l_Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0(x_32, x_10, x_11, x_12, x_13, x_14, x_15); +x_35 = l_Lean_Elab_logException___at___00Lean_Elab_Tactic_elabConfig_spec__0(x_31, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; lean_object* x_37; @@ -7027,7 +7027,7 @@ lean_dec_ref(x_3); lean_dec_ref(x_2); lean_dec(x_1); x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 0, x_31); return x_41; } } @@ -7039,15 +7039,15 @@ if (x_45 == 0) { uint8_t x_46; x_46 = l_Lean_Exception_isRuntime(x_43); -x_31 = lean_box(0); -x_32 = x_43; +x_31 = x_43; +x_32 = lean_box(0); x_33 = x_46; goto block_42; } else { -x_31 = lean_box(0); -x_32 = x_43; +x_31 = x_43; +x_32 = lean_box(0); x_33 = x_45; goto block_42; } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Do/ProofMode/Pure.c b/stage0/stdlib/Lean/Elab/Tactic/Do/ProofMode/Pure.c index b17b4b280de4..132f78a55b71 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Do/ProofMode/Pure.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Do/ProofMode/Pure.c @@ -5642,7 +5642,7 @@ return x_9; { if (x_13 == 0) { -lean_dec_ref(x_12); +lean_dec_ref(x_11); x_7 = lean_box(0); goto block_10; } @@ -5650,7 +5650,7 @@ else { lean_object* x_14; x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 0, x_11); return x_14; } } @@ -5751,15 +5751,15 @@ if (x_42 == 0) { uint8_t x_43; x_43 = l_Lean_Exception_isRuntime(x_41); -x_11 = lean_box(0); -x_12 = x_41; +x_11 = x_41; +x_12 = lean_box(0); x_13 = x_43; goto block_15; } else { -x_11 = lean_box(0); -x_12 = x_41; +x_11 = x_41; +x_12 = lean_box(0); x_13 = x_42; goto block_15; } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Grind/Have.c b/stage0/stdlib/Lean/Elab/Tactic/Grind/Have.c index 781f341f2ddf..8e8cf6142f87 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind/Have.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind/Have.c @@ -2403,13 +2403,13 @@ goto block_272; lean_object* x_25; uint8_t x_26; lean_object* x_27; x_25 = lean_box(0); x_26 = 0; +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -lean_inc(x_19); -lean_inc_ref(x_17); -x_27 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm(x_23, x_25, x_26, x_17, x_19, x_16, x_18, x_14, x_21, x_20, x_15); +lean_inc_ref(x_18); +x_27 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm(x_16, x_25, x_26, x_18, x_15, x_19, x_23, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; @@ -2421,15 +2421,15 @@ x_29 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_29, 0, x_28); x_30 = 0; x_31 = lean_box(0); -lean_inc_ref(x_14); -x_32 = l_Lean_Meta_mkFreshExprMVar(x_29, x_30, x_31, x_14, x_21, x_20, x_15); +lean_inc_ref(x_22); +x_32 = l_Lean_Meta_mkFreshExprMVar(x_29, x_30, x_31, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); lean_dec_ref(x_32); -x_34 = l_Lean_Elab_Tactic_Grind_getGoals___redArg(x_19); +x_34 = l_Lean_Elab_Tactic_Grind_getGoals___redArg(x_15); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; @@ -2442,13 +2442,13 @@ lean_object* x_36; lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_19); -lean_dec_ref(x_17); -x_36 = l_Lean_Elab_Tactic_Grind_throwNoGoalsToBeSolved___redArg(x_14, x_21, x_20, x_15); +lean_dec_ref(x_18); lean_dec(x_15); -lean_dec_ref(x_20); -lean_dec(x_21); -lean_dec_ref(x_14); +x_36 = l_Lean_Elab_Tactic_Grind_throwNoGoalsToBeSolved___redArg(x_22, x_17, x_21, x_20); +lean_dec(x_20); +lean_dec_ref(x_21); +lean_dec(x_17); +lean_dec_ref(x_22); return x_36; } else @@ -2503,19 +2503,19 @@ lean_ctor_set(x_38, 0, x_59); x_60 = lean_box(0); lean_inc_ref(x_38); lean_ctor_set(x_35, 1, x_60); -x_61 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_35, x_19); +x_61 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_35, x_15); if (lean_obj_tag(x_61) == 0) { lean_object* x_62; lean_object* x_63; lean_dec_ref(x_61); x_62 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_solve___boxed), 9, 1); lean_closure_set(x_62, 0, x_38); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -lean_inc_ref(x_17); -x_63 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_62, x_17, x_19, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +lean_inc_ref(x_18); +x_63 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_62, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_63) == 0) { lean_object* x_64; @@ -2525,15 +2525,15 @@ lean_dec_ref(x_63); if (lean_obj_tag(x_64) == 0) { lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_21); +x_65 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_17); x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); lean_dec_ref(x_65); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_67 = l_Lean_MVarId_assert(x_41, x_24, x_28, x_66, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_67 = l_Lean_MVarId_assert(x_41, x_24, x_28, x_66, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_67) == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; @@ -2567,7 +2567,7 @@ lean_closure_set(x_70, 2, x_68); x_71 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_withCheapCasesOnly___boxed), 10, 2); lean_closure_set(x_71, 0, lean_box(0)); lean_closure_set(x_71, 1, x_70); -x_72 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_71, x_17, x_19, x_14, x_21, x_20, x_15); +x_72 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_71, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_72) == 0) { lean_object* x_73; uint8_t x_74; @@ -2582,8 +2582,8 @@ x_75 = lean_ctor_get(x_73, 1); lean_dec(x_75); lean_ctor_set_tag(x_73, 1); lean_ctor_set(x_73, 1, x_40); -x_76 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_73, x_19); -lean_dec(x_19); +x_76 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_73, x_15); +lean_dec(x_15); return x_76; } else @@ -2595,8 +2595,8 @@ lean_dec(x_73); x_78 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_78, 0, x_77); lean_ctor_set(x_78, 1, x_40); -x_79 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_78, x_19); -lean_dec(x_19); +x_79 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_78, x_15); +lean_dec(x_15); return x_79; } } @@ -2604,7 +2604,7 @@ else { uint8_t x_80; lean_dec(x_40); -lean_dec(x_19); +lean_dec(x_15); x_80 = !lean_is_exclusive(x_72); if (x_80 == 0) { @@ -2642,12 +2642,12 @@ lean_dec_ref(x_44); lean_dec_ref(x_43); lean_dec_ref(x_42); lean_dec(x_40); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_83 = !lean_is_exclusive(x_67); if (x_83 == 0) { @@ -2689,28 +2689,28 @@ lean_dec(x_40); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -x_86 = lean_ctor_get(x_17, 3); +x_86 = lean_ctor_get(x_18, 3); lean_inc_ref(x_86); x_87 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkResult___boxed), 10, 2); lean_closure_set(x_87, 0, x_86); lean_closure_set(x_87, 1, x_64); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_88 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_87, x_17, x_19, x_14, x_21, x_20, x_15); -lean_dec(x_19); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_88 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_87, x_18, x_15, x_22, x_17, x_21, x_20); +lean_dec(x_15); if (lean_obj_tag(x_88) == 0) { lean_object* x_89; lean_object* x_90; x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); lean_dec_ref(x_88); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_90 = l_Lean_Meta_Grind_Result_toMessageData(x_89, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_90 = l_Lean_Meta_Grind_Result_toMessageData(x_89, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_90) == 0) { lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; @@ -2721,20 +2721,20 @@ x_92 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHav x_93 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_93, 0, x_92); lean_ctor_set(x_93, 1, x_91); -x_94 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_93, x_14, x_21, x_20, x_15); -lean_dec(x_15); -lean_dec_ref(x_20); -lean_dec(x_21); -lean_dec_ref(x_14); +x_94 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_93, x_22, x_17, x_21, x_20); +lean_dec(x_20); +lean_dec_ref(x_21); +lean_dec(x_17); +lean_dec_ref(x_22); return x_94; } else { uint8_t x_95; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_95 = !lean_is_exclusive(x_90); if (x_95 == 0) { @@ -2755,10 +2755,10 @@ return x_97; else { uint8_t x_98; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_98 = !lean_is_exclusive(x_88); if (x_98 == 0) { @@ -2801,12 +2801,12 @@ lean_dec(x_40); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_101 = !lean_is_exclusive(x_63); if (x_101 == 0) { @@ -2848,12 +2848,12 @@ lean_dec(x_40); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); return x_61; } } @@ -2937,19 +2937,19 @@ x_125 = lean_box(0); lean_inc_ref(x_124); lean_ctor_set(x_35, 1, x_125); lean_ctor_set(x_35, 0, x_124); -x_126 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_35, x_19); +x_126 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_35, x_15); if (lean_obj_tag(x_126) == 0) { lean_object* x_127; lean_object* x_128; lean_dec_ref(x_126); x_127 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_solve___boxed), 9, 1); lean_closure_set(x_127, 0, x_124); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -lean_inc_ref(x_17); -x_128 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_127, x_17, x_19, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +lean_inc_ref(x_18); +x_128 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_127, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_128) == 0) { lean_object* x_129; @@ -2959,15 +2959,15 @@ lean_dec_ref(x_128); if (lean_obj_tag(x_129) == 0) { lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_21); +x_130 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_17); x_131 = lean_ctor_get(x_130, 0); lean_inc(x_131); lean_dec_ref(x_130); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_132 = l_Lean_MVarId_assert(x_105, x_24, x_28, x_131, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_132 = l_Lean_MVarId_assert(x_105, x_24, x_28, x_131, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_132) == 0) { lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; @@ -3001,7 +3001,7 @@ lean_closure_set(x_135, 2, x_133); x_136 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_withCheapCasesOnly___boxed), 10, 2); lean_closure_set(x_136, 0, lean_box(0)); lean_closure_set(x_136, 1, x_135); -x_137 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_136, x_17, x_19, x_14, x_21, x_20, x_15); +x_137 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_136, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_137) == 0) { lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; @@ -3026,15 +3026,15 @@ if (lean_is_scalar(x_140)) { } lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_104); -x_142 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_141, x_19); -lean_dec(x_19); +x_142 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_141, x_15); +lean_dec(x_15); return x_142; } else { lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_dec(x_104); -lean_dec(x_19); +lean_dec(x_15); x_143 = lean_ctor_get(x_137, 0); lean_inc(x_143); if (lean_is_exclusive(x_137)) { @@ -3073,12 +3073,12 @@ lean_dec_ref(x_108); lean_dec_ref(x_107); lean_dec_ref(x_106); lean_dec(x_104); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_146 = lean_ctor_get(x_132, 0); lean_inc(x_146); if (lean_is_exclusive(x_132)) { @@ -3121,28 +3121,28 @@ lean_dec(x_104); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -x_149 = lean_ctor_get(x_17, 3); +x_149 = lean_ctor_get(x_18, 3); lean_inc_ref(x_149); x_150 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkResult___boxed), 10, 2); lean_closure_set(x_150, 0, x_149); lean_closure_set(x_150, 1, x_129); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_151 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_150, x_17, x_19, x_14, x_21, x_20, x_15); -lean_dec(x_19); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_151 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_150, x_18, x_15, x_22, x_17, x_21, x_20); +lean_dec(x_15); if (lean_obj_tag(x_151) == 0) { lean_object* x_152; lean_object* x_153; x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); lean_dec_ref(x_151); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_153 = l_Lean_Meta_Grind_Result_toMessageData(x_152, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_153 = l_Lean_Meta_Grind_Result_toMessageData(x_152, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_153) == 0) { lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; @@ -3153,20 +3153,20 @@ x_155 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHa x_156 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_156, 0, x_155); lean_ctor_set(x_156, 1, x_154); -x_157 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_156, x_14, x_21, x_20, x_15); -lean_dec(x_15); -lean_dec_ref(x_20); -lean_dec(x_21); -lean_dec_ref(x_14); +x_157 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_156, x_22, x_17, x_21, x_20); +lean_dec(x_20); +lean_dec_ref(x_21); +lean_dec(x_17); +lean_dec_ref(x_22); return x_157; } else { lean_object* x_158; lean_object* x_159; lean_object* x_160; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_158 = lean_ctor_get(x_153, 0); lean_inc(x_158); if (lean_is_exclusive(x_153)) { @@ -3188,10 +3188,10 @@ return x_160; else { lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_161 = lean_ctor_get(x_151, 0); lean_inc(x_161); if (lean_is_exclusive(x_151)) { @@ -3235,12 +3235,12 @@ lean_dec(x_104); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_164 = lean_ctor_get(x_128, 0); lean_inc(x_164); if (lean_is_exclusive(x_128)) { @@ -3283,12 +3283,12 @@ lean_dec(x_104); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); return x_126; } } @@ -3404,19 +3404,19 @@ lean_inc_ref(x_189); x_191 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_191, 0, x_189); lean_ctor_set(x_191, 1, x_190); -x_192 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_191, x_19); +x_192 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_191, x_15); if (lean_obj_tag(x_192) == 0) { lean_object* x_193; lean_object* x_194; lean_dec_ref(x_192); x_193 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_solve___boxed), 9, 1); lean_closure_set(x_193, 0, x_189); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -lean_inc_ref(x_17); -x_194 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_193, x_17, x_19, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +lean_inc_ref(x_18); +x_194 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_193, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_194) == 0) { lean_object* x_195; @@ -3426,15 +3426,15 @@ lean_dec_ref(x_194); if (lean_obj_tag(x_195) == 0) { lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_196 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_21); +x_196 = l_Lean_instantiateMVars___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_elabTerm_go_spec__0___redArg(x_33, x_17); x_197 = lean_ctor_get(x_196, 0); lean_inc(x_197); lean_dec_ref(x_196); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_198 = l_Lean_MVarId_assert(x_169, x_24, x_28, x_197, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_198 = l_Lean_MVarId_assert(x_169, x_24, x_28, x_197, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_198) == 0) { lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; @@ -3468,7 +3468,7 @@ lean_closure_set(x_201, 2, x_199); x_202 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_withCheapCasesOnly___boxed), 10, 2); lean_closure_set(x_202, 0, lean_box(0)); lean_closure_set(x_202, 1, x_201); -x_203 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_202, x_17, x_19, x_14, x_21, x_20, x_15); +x_203 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_202, x_18, x_15, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_203) == 0) { lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; @@ -3493,15 +3493,15 @@ if (lean_is_scalar(x_206)) { } lean_ctor_set(x_207, 0, x_205); lean_ctor_set(x_207, 1, x_168); -x_208 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_207, x_19); -lean_dec(x_19); +x_208 = l_Lean_Elab_Tactic_Grind_setGoals___redArg(x_207, x_15); +lean_dec(x_15); return x_208; } else { lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_dec(x_168); -lean_dec(x_19); +lean_dec(x_15); x_209 = lean_ctor_get(x_203, 0); lean_inc(x_209); if (lean_is_exclusive(x_203)) { @@ -3540,12 +3540,12 @@ lean_dec_ref(x_172); lean_dec_ref(x_171); lean_dec_ref(x_170); lean_dec(x_168); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_212 = lean_ctor_get(x_198, 0); lean_inc(x_212); if (lean_is_exclusive(x_198)) { @@ -3588,28 +3588,28 @@ lean_dec(x_168); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -x_215 = lean_ctor_get(x_17, 3); +x_215 = lean_ctor_get(x_18, 3); lean_inc_ref(x_215); x_216 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkResult___boxed), 10, 2); lean_closure_set(x_216, 0, x_215); lean_closure_set(x_216, 1, x_195); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_217 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_216, x_17, x_19, x_14, x_21, x_20, x_15); -lean_dec(x_19); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_217 = l_Lean_Elab_Tactic_Grind_liftGrindM___redArg(x_216, x_18, x_15, x_22, x_17, x_21, x_20); +lean_dec(x_15); if (lean_obj_tag(x_217) == 0) { lean_object* x_218; lean_object* x_219; x_218 = lean_ctor_get(x_217, 0); lean_inc(x_218); lean_dec_ref(x_217); -lean_inc(x_15); -lean_inc_ref(x_20); -lean_inc(x_21); -lean_inc_ref(x_14); -x_219 = l_Lean_Meta_Grind_Result_toMessageData(x_218, x_14, x_21, x_20, x_15); +lean_inc(x_20); +lean_inc_ref(x_21); +lean_inc(x_17); +lean_inc_ref(x_22); +x_219 = l_Lean_Meta_Grind_Result_toMessageData(x_218, x_22, x_17, x_21, x_20); if (lean_obj_tag(x_219) == 0) { lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; @@ -3620,20 +3620,20 @@ x_221 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHa x_222 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_222, 0, x_221); lean_ctor_set(x_222, 1, x_220); -x_223 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_222, x_14, x_21, x_20, x_15); -lean_dec(x_15); -lean_dec_ref(x_20); -lean_dec(x_21); -lean_dec_ref(x_14); +x_223 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHave_spec__2___redArg(x_222, x_22, x_17, x_21, x_20); +lean_dec(x_20); +lean_dec_ref(x_21); +lean_dec(x_17); +lean_dec_ref(x_22); return x_223; } else { lean_object* x_224; lean_object* x_225; lean_object* x_226; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_224 = lean_ctor_get(x_219, 0); lean_inc(x_224); if (lean_is_exclusive(x_219)) { @@ -3655,10 +3655,10 @@ return x_226; else { lean_object* x_227; lean_object* x_228; lean_object* x_229; -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_15); -lean_dec_ref(x_14); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec(x_17); x_227 = lean_ctor_get(x_217, 0); lean_inc(x_227); if (lean_is_exclusive(x_217)) { @@ -3702,12 +3702,12 @@ lean_dec(x_168); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_230 = lean_ctor_get(x_194, 0); lean_inc(x_230); if (lean_is_exclusive(x_194)) { @@ -3750,12 +3750,12 @@ lean_dec(x_168); lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); return x_192; } } @@ -3767,12 +3767,12 @@ uint8_t x_233; lean_dec(x_33); lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_233 = !lean_is_exclusive(x_34); if (x_233 == 0) { @@ -3795,12 +3795,12 @@ else uint8_t x_236; lean_dec(x_28); lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_236 = !lean_is_exclusive(x_32); if (x_236 == 0) { @@ -3822,12 +3822,12 @@ else { uint8_t x_239; lean_dec(x_24); -lean_dec(x_21); -lean_dec_ref(x_20); -lean_dec(x_19); -lean_dec_ref(x_17); +lean_dec_ref(x_22); +lean_dec_ref(x_21); +lean_dec(x_20); +lean_dec_ref(x_18); +lean_dec(x_17); lean_dec(x_15); -lean_dec_ref(x_14); x_239 = !lean_is_exclusive(x_27); if (x_239 == 0) { @@ -3853,31 +3853,31 @@ if (x_254 == 0) { lean_object* x_255; x_255 = l_Lean_Meta_Grind_markGrindName(x_253); -x_14 = x_243; +x_14 = lean_box(0); x_15 = x_244; -x_16 = x_246; -x_17 = x_245; -x_18 = x_248; -x_19 = x_247; -x_20 = x_249; -x_21 = x_250; -x_22 = lean_box(0); -x_23 = x_251; +x_16 = x_243; +x_17 = x_246; +x_18 = x_247; +x_19 = x_248; +x_20 = x_250; +x_21 = x_249; +x_22 = x_251; +x_23 = x_252; x_24 = x_255; goto block_242; } else { -x_14 = x_243; +x_14 = lean_box(0); x_15 = x_244; -x_16 = x_246; -x_17 = x_245; -x_18 = x_248; -x_19 = x_247; -x_20 = x_249; -x_21 = x_250; -x_22 = lean_box(0); -x_23 = x_251; +x_16 = x_243; +x_17 = x_246; +x_18 = x_247; +x_19 = x_248; +x_20 = x_250; +x_21 = x_249; +x_22 = x_251; +x_23 = x_252; x_24 = x_253; goto block_242; } @@ -3891,16 +3891,16 @@ if (lean_obj_tag(x_257) == 0) { lean_object* x_269; x_269 = l___private_Lean_Elab_Tactic_Grind_Have_0__Lean_Elab_Tactic_Grind_evalHaveSilent___lam__0___closed__3; -x_243 = x_262; -x_244 = x_265; -x_245 = x_258; -x_246 = x_260; -x_247 = x_259; -x_248 = x_261; +x_243 = x_268; +x_244 = x_259; +x_245 = lean_box(0); +x_246 = x_263; +x_247 = x_258; +x_248 = x_260; x_249 = x_264; -x_250 = x_263; -x_251 = x_268; -x_252 = lean_box(0); +x_250 = x_265; +x_251 = x_262; +x_252 = x_261; x_253 = x_269; goto block_256; } @@ -3912,16 +3912,16 @@ lean_inc(x_270); lean_dec_ref(x_257); x_271 = l_Lean_TSyntax_getId(x_270); lean_dec(x_270); -x_243 = x_262; -x_244 = x_265; -x_245 = x_258; -x_246 = x_260; -x_247 = x_259; -x_248 = x_261; +x_243 = x_268; +x_244 = x_259; +x_245 = lean_box(0); +x_246 = x_263; +x_247 = x_258; +x_248 = x_260; x_249 = x_264; -x_250 = x_263; -x_251 = x_268; -x_252 = lean_box(0); +x_250 = x_265; +x_251 = x_262; +x_252 = x_261; x_253 = x_271; goto block_256; } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Grind/Main.c b/stage0/stdlib/Lean/Elab/Tactic/Grind/Main.c index 79228f727795..5c5aa3bde6a6 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind/Main.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind/Main.c @@ -11768,17 +11768,17 @@ goto block_30; block_58: { lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = l_Lean_PersistentArray_push___redArg(x_46, x_45); -x_55 = l_Lean_PersistentArray_push___redArg(x_54, x_44); +x_54 = l_Lean_PersistentArray_push___redArg(x_44, x_47); +x_55 = l_Lean_PersistentArray_push___redArg(x_54, x_52); x_56 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_47); +lean_ctor_set(x_56, 0, x_48); +lean_ctor_set(x_56, 1, x_43); lean_ctor_set(x_56, 2, x_50); -lean_ctor_set(x_56, 3, x_49); -lean_ctor_set(x_56, 4, x_52); +lean_ctor_set(x_56, 3, x_46); +lean_ctor_set(x_56, 4, x_49); lean_ctor_set(x_56, 5, x_55); -lean_ctor_set(x_56, 6, x_48); -lean_ctor_set(x_56, 7, x_43); +lean_ctor_set(x_56, 6, x_51); +lean_ctor_set(x_56, 7, x_45); x_57 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_57, 0, x_56); return x_57; @@ -12034,13 +12034,13 @@ return x_117; { lean_object* x_124; lean_object* x_125; x_124 = lean_ctor_get(x_1, 3); -lean_inc(x_119); -lean_inc_ref(x_121); -lean_inc(x_120); +lean_inc(x_122); +lean_inc_ref(x_120); +lean_inc(x_121); lean_inc_ref(x_123); lean_inc_ref(x_124); lean_inc(x_3); -x_125 = l_Lean_Meta_Grind_mkEMatchTheoremForDecl(x_3, x_4, x_124, x_59, x_5, x_123, x_120, x_121, x_119); +x_125 = l_Lean_Meta_Grind_mkEMatchTheoremForDecl(x_3, x_4, x_124, x_59, x_5, x_123, x_121, x_120, x_122); if (lean_obj_tag(x_125) == 0) { lean_object* x_126; @@ -12049,9 +12049,9 @@ lean_inc(x_126); lean_dec_ref(x_125); x_31 = x_126; x_32 = x_123; -x_33 = x_120; -x_34 = x_121; -x_35 = x_119; +x_33 = x_121; +x_34 = x_120; +x_35 = x_122; x_36 = lean_box(0); goto block_42; } @@ -12059,9 +12059,9 @@ else { uint8_t x_127; lean_dec_ref(x_123); -lean_dec_ref(x_121); -lean_dec(x_120); -lean_dec(x_119); +lean_dec(x_122); +lean_dec(x_121); +lean_dec_ref(x_120); lean_dec(x_3); lean_dec_ref(x_1); x_127 = !lean_is_exclusive(x_125); @@ -12086,10 +12086,10 @@ return x_129; if (x_6 == 0) { lean_dec(x_2); -x_119 = x_134; -x_120 = x_132; -x_121 = x_133; -x_122 = lean_box(0); +x_119 = lean_box(0); +x_120 = x_133; +x_121 = x_132; +x_122 = x_134; x_123 = x_131; goto block_130; } @@ -12154,10 +12154,10 @@ return x_144; else { lean_dec(x_2); -x_119 = x_134; -x_120 = x_132; -x_121 = x_133; -x_122 = lean_box(0); +x_119 = lean_box(0); +x_120 = x_133; +x_121 = x_132; +x_122 = x_134; x_123 = x_131; goto block_130; } @@ -12166,23 +12166,23 @@ goto block_130; block_155: { lean_object* x_151; -x_151 = l___private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_ensureNoMinIndexable(x_5, x_150, x_147, x_148, x_146); +x_151 = l___private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_ensureNoMinIndexable(x_5, x_147, x_146, x_149, x_150); if (lean_obj_tag(x_151) == 0) { lean_dec_ref(x_151); -x_131 = x_150; -x_132 = x_147; -x_133 = x_148; -x_134 = x_146; +x_131 = x_147; +x_132 = x_146; +x_133 = x_149; +x_134 = x_150; x_135 = lean_box(0); goto block_145; } else { uint8_t x_152; -lean_dec_ref(x_150); -lean_dec_ref(x_148); -lean_dec(x_147); +lean_dec(x_150); +lean_dec_ref(x_149); +lean_dec_ref(x_147); lean_dec(x_146); lean_dec(x_4); lean_dec(x_3); @@ -12274,16 +12274,16 @@ lean_dec(x_3); x_176 = lean_ctor_get(x_175, 0); lean_inc(x_176); lean_dec_ref(x_175); -x_43 = x_171; -x_44 = x_176; -x_45 = x_173; -x_46 = x_169; -x_47 = x_165; -x_48 = x_170; -x_49 = x_167; +x_43 = x_165; +x_44 = x_169; +x_45 = x_171; +x_46 = x_167; +x_47 = x_173; +x_48 = x_164; +x_49 = x_168; x_50 = x_166; -x_51 = x_164; -x_52 = x_168; +x_51 = x_170; +x_52 = x_176; x_53 = lean_box(0); goto block_58; } @@ -12305,16 +12305,16 @@ lean_dec_ref(x_158); lean_dec(x_157); lean_dec_ref(x_156); lean_dec(x_3); -x_43 = x_171; -x_44 = x_177; -x_45 = x_173; -x_46 = x_169; -x_47 = x_165; -x_48 = x_170; -x_49 = x_167; +x_43 = x_165; +x_44 = x_169; +x_45 = x_171; +x_46 = x_167; +x_47 = x_173; +x_48 = x_164; +x_49 = x_168; x_50 = x_166; -x_51 = x_164; -x_52 = x_168; +x_51 = x_170; +x_52 = x_177; x_53 = lean_box(0); goto block_58; } @@ -12333,16 +12333,16 @@ lean_dec_ref(x_158); lean_dec(x_157); lean_dec_ref(x_156); lean_dec(x_3); -x_43 = x_171; -x_44 = x_177; -x_45 = x_173; -x_46 = x_169; -x_47 = x_165; -x_48 = x_170; -x_49 = x_167; +x_43 = x_165; +x_44 = x_169; +x_45 = x_171; +x_46 = x_167; +x_47 = x_173; +x_48 = x_164; +x_49 = x_168; x_50 = x_166; -x_51 = x_164; -x_52 = x_168; +x_51 = x_170; +x_52 = x_177; x_53 = lean_box(0); goto block_58; } @@ -12355,16 +12355,16 @@ lean_dec(x_159); lean_dec(x_157); lean_dec_ref(x_156); lean_dec_ref(x_184); -x_43 = x_171; -x_44 = x_177; -x_45 = x_173; -x_46 = x_169; -x_47 = x_165; -x_48 = x_170; -x_49 = x_167; +x_43 = x_165; +x_44 = x_169; +x_45 = x_171; +x_46 = x_167; +x_47 = x_173; +x_48 = x_164; +x_49 = x_168; x_50 = x_166; -x_51 = x_164; -x_52 = x_168; +x_51 = x_170; +x_52 = x_177; x_53 = lean_box(0); goto block_58; } @@ -12529,16 +12529,16 @@ lean_dec(x_3); x_209 = lean_ctor_get(x_208, 0); lean_inc(x_209); lean_dec_ref(x_208); -x_43 = x_203; -x_44 = x_209; -x_45 = x_206; -x_46 = x_201; -x_47 = x_197; -x_48 = x_202; -x_49 = x_199; +x_43 = x_197; +x_44 = x_201; +x_45 = x_203; +x_46 = x_199; +x_47 = x_206; +x_48 = x_196; +x_49 = x_200; x_50 = x_198; -x_51 = x_196; -x_52 = x_200; +x_51 = x_202; +x_52 = x_209; x_53 = lean_box(0); goto block_58; } @@ -12560,16 +12560,16 @@ lean_dec_ref(x_158); lean_dec(x_157); lean_dec_ref(x_156); lean_dec(x_3); -x_43 = x_203; -x_44 = x_210; -x_45 = x_206; -x_46 = x_201; -x_47 = x_197; -x_48 = x_202; -x_49 = x_199; +x_43 = x_197; +x_44 = x_201; +x_45 = x_203; +x_46 = x_199; +x_47 = x_206; +x_48 = x_196; +x_49 = x_200; x_50 = x_198; -x_51 = x_196; -x_52 = x_200; +x_51 = x_202; +x_52 = x_210; x_53 = lean_box(0); goto block_58; } @@ -12588,16 +12588,16 @@ lean_dec_ref(x_158); lean_dec(x_157); lean_dec_ref(x_156); lean_dec(x_3); -x_43 = x_203; -x_44 = x_210; -x_45 = x_206; -x_46 = x_201; -x_47 = x_197; -x_48 = x_202; -x_49 = x_199; +x_43 = x_197; +x_44 = x_201; +x_45 = x_203; +x_46 = x_199; +x_47 = x_206; +x_48 = x_196; +x_49 = x_200; x_50 = x_198; -x_51 = x_196; -x_52 = x_200; +x_51 = x_202; +x_52 = x_210; x_53 = lean_box(0); goto block_58; } @@ -12610,16 +12610,16 @@ lean_dec(x_159); lean_dec(x_157); lean_dec_ref(x_156); lean_dec_ref(x_217); -x_43 = x_203; -x_44 = x_210; -x_45 = x_206; -x_46 = x_201; -x_47 = x_197; -x_48 = x_202; -x_49 = x_199; +x_43 = x_197; +x_44 = x_201; +x_45 = x_203; +x_46 = x_199; +x_47 = x_206; +x_48 = x_196; +x_49 = x_200; x_50 = x_198; -x_51 = x_196; -x_52 = x_200; +x_51 = x_202; +x_52 = x_210; x_53 = lean_box(0); goto block_58; } @@ -12728,20 +12728,20 @@ else switch (lean_obj_tag(x_4)) { case 0: { -x_146 = x_159; -x_147 = x_157; -x_148 = x_158; -x_149 = lean_box(0); -x_150 = x_156; +x_146 = x_157; +x_147 = x_156; +x_148 = lean_box(0); +x_149 = x_158; +x_150 = x_159; goto block_155; } case 1: { -x_146 = x_159; -x_147 = x_157; -x_148 = x_158; -x_149 = lean_box(0); -x_150 = x_156; +x_146 = x_157; +x_147 = x_156; +x_148 = lean_box(0); +x_149 = x_158; +x_150 = x_159; goto block_155; } default: @@ -16899,14 +16899,14 @@ lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec_ref(x_21); -lean_inc(x_12); +lean_inc(x_13); x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_12); +lean_ctor_set(x_23, 0, x_13); x_24 = l_Lean_Meta_Grind_Theorems_find___redArg(x_22, x_23); lean_dec_ref(x_23); x_25 = lean_box(0); -x_26 = l_List_filterTR_loop___at___00__private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_processParam_spec__1(x_13, x_24, x_25); -lean_dec(x_13); +x_26 = l_List_filterTR_loop___at___00__private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_processParam_spec__1(x_12, x_24, x_25); +lean_dec(x_12); x_27 = l_List_isEmpty___redArg(x_26); if (x_27 == 0) { @@ -16915,7 +16915,7 @@ lean_dec(x_18); lean_dec_ref(x_17); lean_dec(x_16); lean_dec_ref(x_15); -lean_dec(x_12); +lean_dec(x_13); x_28 = l_List_forIn_x27_loop___at___00__private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_processParam_spec__2___redArg(x_26, x_14); return x_28; } @@ -16926,7 +16926,7 @@ lean_dec(x_26); lean_dec_ref(x_14); x_29 = l___private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindParams_processParam___closed__1; x_30 = 0; -x_31 = l_Lean_MessageData_ofConstName(x_12, x_30); +x_31 = l_Lean_MessageData_ofConstName(x_13, x_30); x_32 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_32, 0, x_29); lean_ctor_set(x_32, 1, x_31); @@ -17440,8 +17440,8 @@ lean_dec_ref(x_154); if (lean_obj_tag(x_155) == 0) { lean_dec_ref(x_155); -x_12 = x_132; -x_13 = x_136; +x_12 = x_136; +x_13 = x_132; x_14 = x_2; x_15 = x_7; x_16 = x_8; @@ -17478,8 +17478,8 @@ return x_158; } else { -x_12 = x_132; -x_13 = x_136; +x_12 = x_136; +x_13 = x_132; x_14 = x_2; x_15 = x_7; x_16 = x_8; @@ -20352,18 +20352,18 @@ if (x_37 == 0) uint8_t x_38; x_38 = l_Lean_Exception_isRuntime(x_36); lean_dec(x_36); -x_16 = x_34; +x_16 = lean_box(0); x_17 = x_24; -x_18 = lean_box(0); +x_18 = x_34; x_19 = x_38; goto block_20; } else { lean_dec(x_36); -x_16 = x_34; +x_16 = lean_box(0); x_17 = x_24; -x_18 = lean_box(0); +x_18 = x_34; x_19 = x_37; goto block_20; } @@ -20383,7 +20383,7 @@ goto _start; { if (x_19 == 0) { -lean_dec_ref(x_16); +lean_dec_ref(x_18); x_10 = x_17; x_11 = lean_box(0); goto block_15; @@ -20395,7 +20395,7 @@ lean_dec(x_8); lean_dec_ref(x_7); lean_dec(x_6); lean_dec_ref(x_5); -return x_16; +return x_18; } } } @@ -22872,41 +22872,41 @@ return x_149; block_56: { lean_object* x_51; -lean_inc(x_50); -lean_inc_ref(x_47); -lean_inc(x_41); +lean_inc(x_48); lean_inc_ref(x_45); -x_51 = l_Lean_Meta_instantiateMVarsProfiling(x_40, x_45, x_41, x_47, x_50); +lean_inc(x_46); +lean_inc_ref(x_50); +x_51 = l_Lean_Meta_instantiateMVarsProfiling(x_40, x_50, x_46, x_45, x_48); if (lean_obj_tag(x_51) == 0) { lean_object* x_52; x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); lean_dec_ref(x_51); -x_16 = x_43; +x_16 = x_44; x_17 = x_52; -x_18 = x_49; -x_19 = x_48; -x_20 = x_42; -x_21 = x_44; -x_22 = x_45; -x_23 = x_41; -x_24 = x_47; -x_25 = x_50; +x_18 = x_41; +x_19 = x_47; +x_20 = x_43; +x_21 = x_42; +x_22 = x_50; +x_23 = x_46; +x_24 = x_45; +x_25 = x_48; x_26 = lean_box(0); goto block_33; } else { uint8_t x_53; -lean_dec(x_50); -lean_dec_ref(x_49); -lean_dec_ref(x_47); +lean_dec_ref(x_50); +lean_dec(x_48); +lean_dec(x_46); 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_41); +lean_dec(x_42); +lean_dec_ref(x_41); lean_dec(x_4); x_53 = !lean_is_exclusive(x_51); if (x_53 == 0) @@ -22944,16 +22944,16 @@ if (x_69 == 0) lean_dec_ref(x_68); lean_dec_ref(x_57); lean_dec(x_37); -x_41 = x_64; -x_42 = x_61; -x_43 = x_58; -x_44 = x_62; -x_45 = x_63; -x_46 = lean_box(0); -x_47 = x_65; -x_48 = x_60; -x_49 = x_59; -x_50 = x_66; +x_41 = x_59; +x_42 = x_62; +x_43 = x_61; +x_44 = x_58; +x_45 = x_65; +x_46 = x_64; +x_47 = x_60; +x_48 = x_66; +x_49 = lean_box(0); +x_50 = x_63; goto block_56; } else @@ -23208,16 +23208,16 @@ else lean_dec_ref(x_68); lean_dec_ref(x_57); lean_dec(x_37); -x_41 = x_64; -x_42 = x_61; -x_43 = x_58; -x_44 = x_62; -x_45 = x_63; -x_46 = lean_box(0); -x_47 = x_65; -x_48 = x_60; -x_49 = x_59; -x_50 = x_66; +x_41 = x_59; +x_42 = x_62; +x_43 = x_61; +x_44 = x_58; +x_45 = x_65; +x_46 = x_64; +x_47 = x_60; +x_48 = x_66; +x_49 = lean_box(0); +x_50 = x_63; goto block_56; } } @@ -23845,7 +23845,7 @@ return x_16; LEAN_EXPORT lean_object* l_Lean_logAt___at___00Lean_logWarningAt___at___00Lean_Elab_Tactic_evalGrindCore_spec__0_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) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; 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; 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; uint8_t x_89; lean_object* x_90; lean_object* x_91; uint8_t 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; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_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; uint8_t 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; 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; 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) @@ -23880,15 +23880,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_13); x_27 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_10); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_12); lean_ctor_set(x_27, 2, x_15); -lean_ctor_set(x_27, 3, x_12); +lean_ctor_set(x_27, 3, x_11); lean_ctor_set(x_27, 4, x_26); lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_16); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_13); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_14); 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); @@ -23925,15 +23925,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_13); x_43 = lean_alloc_ctor(0, 5, 3); -lean_ctor_set(x_43, 0, x_14); -lean_ctor_set(x_43, 1, x_10); +lean_ctor_set(x_43, 0, x_10); +lean_ctor_set(x_43, 1, x_12); lean_ctor_set(x_43, 2, x_15); -lean_ctor_set(x_43, 3, x_12); +lean_ctor_set(x_43, 3, x_11); lean_ctor_set(x_43, 4, x_42); lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_16); -lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_13); +lean_ctor_set_uint8(x_43, sizeof(void*)*5 + 1, x_14); 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); @@ -23963,23 +23963,23 @@ 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_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_addTrace___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_elabGrindConfig_spec__0_spec__0_spec__4___redArg___closed__1; -if (x_51 == 0) +if (x_54 == 0) { lean_free_object(x_59); lean_dec_ref(x_50); -x_10 = x_62; -x_11 = x_61; -x_12 = x_65; -x_13 = x_55; -x_14 = x_54; +x_10 = x_51; +x_11 = x_65; +x_12 = x_62; +x_13 = x_61; +x_14 = x_53; x_15 = x_64; x_16 = x_56; x_17 = x_7; @@ -23998,7 +23998,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); @@ -24007,11 +24007,11 @@ return x_59; else { lean_free_object(x_59); -x_10 = x_62; -x_11 = x_61; -x_12 = x_65; -x_13 = x_55; -x_14 = x_54; +x_10 = x_51; +x_11 = x_65; +x_12 = x_62; +x_13 = x_61; +x_14 = x_53; x_15 = x_64; x_16 = x_56; x_17 = x_7; @@ -24027,22 +24027,22 @@ 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_52); -lean_dec(x_52); -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_addTrace___at___00__private_Lean_ExtraModUses_0__Lean_recordExtraModUseCore___at___00Lean_recordExtraModUseFromDecl___at___00Lean_Elab_Tactic_elabGrindConfig_spec__0_spec__0_spec__4___redArg___closed__1; -if (x_51 == 0) +if (x_54 == 0) { lean_dec_ref(x_50); -x_10 = x_69; -x_11 = x_68; -x_12 = x_72; -x_13 = x_55; -x_14 = x_54; +x_10 = x_51; +x_11 = x_72; +x_12 = x_69; +x_13 = x_68; +x_14 = x_53; x_15 = x_71; x_16 = x_56; x_17 = x_7; @@ -24061,7 +24061,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); @@ -24070,11 +24070,11 @@ return x_75; } else { -x_10 = x_69; -x_11 = x_68; -x_12 = x_72; -x_13 = x_55; -x_14 = x_54; +x_10 = x_51; +x_11 = x_72; +x_12 = x_69; +x_13 = x_68; +x_14 = x_53; x_15 = x_71; x_16 = x_56; x_17 = x_7; @@ -24088,17 +24088,17 @@ goto block_49; block_87: { lean_object* x_85; -x_85 = l_Lean_Syntax_getTailPos_x3f(x_80, x_83); -lean_dec(x_80); +x_85 = l_Lean_Syntax_getTailPos_x3f(x_82, x_83); +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_84; -x_53 = x_79; -x_54 = x_82; -x_55 = x_81; +x_52 = x_79; +x_53 = x_80; +x_54 = x_81; +x_55 = x_84; x_56 = x_83; x_57 = x_84; goto block_76; @@ -24111,10 +24111,10 @@ 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_82; -x_55 = x_81; +x_52 = x_79; +x_53 = x_80; +x_54 = x_81; +x_55 = x_84; x_56 = x_83; x_57 = x_86; goto block_76; @@ -24123,9 +24123,9 @@ 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_96 = l_Lean_Syntax_getPos_x3f(x_95, x_92); +x_95 = l_Lean_replaceRef(x_1, x_91); +lean_dec(x_91); +x_96 = l_Lean_Syntax_getPos_x3f(x_95, x_93); if (lean_obj_tag(x_96) == 0) { lean_object* x_97; @@ -24133,10 +24133,10 @@ x_97 = lean_unsigned_to_nat(0u); x_77 = x_88; x_78 = x_89; x_79 = x_90; -x_80 = x_95; -x_81 = x_94; -x_82 = x_91; -x_83 = x_92; +x_80 = x_94; +x_81 = x_92; +x_82 = x_95; +x_83 = x_93; x_84 = x_97; goto block_87; } @@ -24149,10 +24149,10 @@ lean_dec_ref(x_96); x_77 = x_88; x_78 = x_89; x_79 = x_90; -x_80 = x_95; -x_81 = x_94; -x_82 = x_91; -x_83 = x_92; +x_80 = x_94; +x_81 = x_92; +x_82 = x_95; +x_83 = x_93; x_84 = x_98; goto block_87; } @@ -24161,23 +24161,23 @@ goto block_87; { if (x_107 == 0) { -x_88 = x_104; +x_88 = x_103; x_89 = x_101; x_90 = x_102; -x_91 = x_103; -x_92 = x_106; -x_93 = x_105; +x_91 = x_104; +x_92 = x_105; +x_93 = x_106; x_94 = x_3; goto block_99; } else { -x_88 = x_104; +x_88 = x_103; x_89 = x_101; x_90 = x_102; -x_91 = x_103; -x_92 = x_106; -x_93 = x_105; +x_91 = x_104; +x_92 = x_105; +x_93 = x_106; x_94 = x_100; goto block_99; } @@ -24202,13 +24202,13 @@ 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_ref(x_111); -x_101 = x_114; +lean_inc_ref(x_110); +x_101 = x_110; x_102 = x_111; -x_103 = x_110; -x_104 = x_117; -x_105 = x_113; +x_103 = x_117; +x_104 = x_113; +x_105 = x_114; x_106 = x_109; x_107 = x_119; goto block_108; @@ -24219,13 +24219,13 @@ 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_Main_0__Lean_Elab_Tactic_warnRedundantEMatchArg_spec__0_spec__0_spec__0___closed__0; x_121 = l_Lean_Option_get___at___00Lean_Elab_addMacroStack___at___00Lean_throwError___at___00Lean_Elab_Tactic_elabGrindConfig_spec__8_spec__8_spec__8(x_112, x_120); lean_inc(x_113); -lean_inc_ref(x_110); lean_inc_ref(x_111); -x_101 = x_114; +lean_inc_ref(x_110); +x_101 = x_110; x_102 = x_111; -x_103 = x_110; -x_104 = x_117; -x_105 = x_113; +x_103 = x_117; +x_104 = x_113; +x_105 = x_114; x_106 = x_109; x_107 = x_121; goto block_108; @@ -25918,8 +25918,8 @@ if (x_43 == 0) lean_dec(x_41); lean_dec_ref(x_39); lean_dec(x_33); -x_23 = x_37; -x_24 = lean_box(0); +x_23 = lean_box(0); +x_24 = x_37; x_25 = x_42; goto block_28; } @@ -25932,8 +25932,8 @@ if (x_44 == 0) lean_dec(x_41); lean_dec_ref(x_39); lean_dec(x_33); -x_23 = x_37; -x_24 = lean_box(0); +x_23 = lean_box(0); +x_24 = x_37; x_25 = x_42; goto block_28; } @@ -25946,8 +25946,8 @@ lean_dec(x_41); x_47 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_unresolveNameGlobal___at___00Lean_unresolveNameGlobalAvoidingLocals___at___00Lean_Elab_Tactic_mkGrindOnly_spec__0_spec__0_spec__3(x_33, x_39, x_45, x_46, x_42); lean_dec_ref(x_39); lean_dec(x_33); -x_23 = x_37; -x_24 = lean_box(0); +x_23 = lean_box(0); +x_24 = x_37; x_25 = x_47; goto block_28; } @@ -26123,7 +26123,7 @@ return x_21; { lean_object* x_26; lean_object* x_27; x_26 = lean_box(0); -x_27 = lean_apply_7(x_23, x_25, x_26, x_5, x_6, x_7, x_8, lean_box(0)); +x_27 = lean_apply_7(x_24, x_25, x_26, x_5, x_6, x_7, x_8, lean_box(0)); return x_27; } } @@ -29640,18 +29640,18 @@ if (lean_is_scalar(x_14)) { lean_ctor_set_tag(x_72, 1); } lean_ctor_set(x_72, 0, x_68); -x_20 = x_56; -x_21 = x_60; -x_22 = x_54; -x_23 = x_55; -x_24 = x_59; -x_25 = x_53; +x_20 = x_53; +x_21 = x_54; +x_22 = x_58; +x_23 = lean_box(0); +x_24 = x_56; +x_25 = x_60; x_26 = x_57; -x_27 = x_52; -x_28 = lean_box(0); -x_29 = x_72; -x_30 = x_58; -x_31 = x_61; +x_27 = x_72; +x_28 = x_55; +x_29 = x_61; +x_30 = x_52; +x_31 = x_59; x_32 = x_49; goto block_47; } @@ -29663,18 +29663,18 @@ lean_object* x_73; lean_dec(x_64); lean_dec(x_14); x_73 = lean_box(0); -x_20 = x_56; -x_21 = x_60; -x_22 = x_54; -x_23 = x_55; -x_24 = x_59; -x_25 = x_53; +x_20 = x_53; +x_21 = x_54; +x_22 = x_58; +x_23 = lean_box(0); +x_24 = x_56; +x_25 = x_60; x_26 = x_57; -x_27 = x_52; -x_28 = lean_box(0); -x_29 = x_73; -x_30 = x_58; -x_31 = x_61; +x_27 = x_73; +x_28 = x_55; +x_29 = x_61; +x_30 = x_52; +x_31 = x_59; x_32 = x_12; goto block_47; } @@ -29755,22 +29755,22 @@ goto block_74; block_47: { lean_object* x_33; +lean_inc(x_29); +lean_inc_ref(x_25); lean_inc(x_31); -lean_inc_ref(x_21); -lean_inc(x_24); -lean_inc_ref(x_30); +lean_inc_ref(x_22); lean_inc(x_26); -lean_inc_ref(x_20); -x_33 = l___private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindConfig_x27___redArg(x_19, x_32, x_22, x_20, x_26, x_30, x_24, x_21, x_31); +lean_inc_ref(x_24); +x_33 = l___private_Lean_Elab_Tactic_Grind_Main_0__Lean_Elab_Tactic_elabGrindConfig_x27___redArg(x_19, x_32, x_21, x_24, x_26, x_22, x_31, x_25, x_29); 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); -x_35 = l_Lean_Elab_Tactic_evalGrindCore(x_1, x_34, x_27, x_25, x_29, x_22, x_23, x_20, x_26, x_30, x_24, x_21, x_31); -lean_dec(x_25); -lean_dec(x_27); +x_35 = l_Lean_Elab_Tactic_evalGrindCore(x_1, x_34, x_30, x_20, x_27, x_21, x_28, x_24, x_26, x_22, x_31, x_25, x_29); +lean_dec(x_20); +lean_dec(x_30); lean_dec(x_1); if (lean_obj_tag(x_35) == 0) { @@ -29819,16 +29819,16 @@ else { uint8_t x_44; lean_dec(x_31); -lean_dec_ref(x_30); +lean_dec(x_30); lean_dec(x_29); +lean_dec(x_28); lean_dec(x_27); lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); +lean_dec_ref(x_25); +lean_dec_ref(x_24); lean_dec_ref(x_22); lean_dec_ref(x_21); -lean_dec_ref(x_20); +lean_dec(x_20); lean_dec(x_1); x_44 = !lean_is_exclusive(x_33); if (x_44 == 0) diff --git a/stage0/stdlib/Lean/Elab/Tactic/RCases.c b/stage0/stdlib/Lean/Elab/Tactic/RCases.c index 181f67ac22f7..2cc76192644a 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/RCases.c +++ b/stage0/stdlib/Lean/Elab/Tactic/RCases.c @@ -3774,10 +3774,10 @@ return x_9; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 1, x_13); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_13); +lean_ctor_set(x_16, 1, x_12); x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -3799,8 +3799,8 @@ if (lean_obj_tag(x_25) == 0) lean_object* x_26; x_26 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__1; x_11 = x_19; -x_12 = x_23; -x_13 = x_24; +x_12 = x_24; +x_13 = x_23; x_14 = x_26; goto block_18; } @@ -3811,8 +3811,8 @@ x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec_ref(x_25); x_11 = x_19; -x_12 = x_23; -x_13 = x_24; +x_12 = x_24; +x_13 = x_23; x_14 = x_27; goto block_18; } @@ -15972,14 +15972,14 @@ goto block_39; block_19: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = l_Array_toSubarray___redArg(x_11, x_12, x_13); +x_14 = l_Array_toSubarray___redArg(x_10, x_12, x_13); x_15 = l_Array_ofSubarray___redArg(x_14); lean_dec_ref(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_8); +lean_ctor_set(x_16, 1, x_11); x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_16); x_18 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -16017,10 +16017,10 @@ x_34 = lean_array_get_size(x_25); x_35 = lean_nat_dec_le(x_32, x_20); if (x_35 == 0) { -x_8 = x_26; +x_8 = x_33; x_9 = lean_box(0); -x_10 = x_33; -x_11 = x_25; +x_10 = x_25; +x_11 = x_26; x_12 = x_32; x_13 = x_34; goto block_19; @@ -16028,10 +16028,10 @@ goto block_19; else { lean_dec(x_32); -x_8 = x_26; +x_8 = x_33; x_9 = lean_box(0); -x_10 = x_33; -x_11 = x_25; +x_10 = x_25; +x_11 = x_26; x_12 = x_20; x_13 = x_34; goto block_19; @@ -16554,10 +16554,10 @@ lean_dec_ref(x_45); x_47 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_RCasesPatt_name_x3f(x_33); if (lean_obj_tag(x_31) == 0) { -x_12 = x_33; +x_12 = x_47; x_13 = lean_box(0); -x_14 = x_47; -x_15 = x_46; +x_14 = x_46; +x_15 = x_33; x_16 = x_44; goto block_23; } @@ -16572,10 +16572,10 @@ x_49 = lean_ctor_get(x_31, 0); x_50 = l_Lean_TSyntax_getId(x_49); lean_dec(x_49); lean_ctor_set(x_31, 0, x_50); -x_12 = x_33; +x_12 = x_47; x_13 = lean_box(0); -x_14 = x_47; -x_15 = x_46; +x_14 = x_46; +x_15 = x_33; x_16 = x_31; goto block_23; } @@ -16589,10 +16589,10 @@ x_52 = l_Lean_TSyntax_getId(x_51); lean_dec(x_51); x_53 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_53, 0, x_52); -x_12 = x_33; +x_12 = x_47; x_13 = lean_box(0); -x_14 = x_47; -x_15 = x_46; +x_14 = x_46; +x_15 = x_33; x_16 = x_53; goto block_23; } @@ -16671,11 +16671,11 @@ return x_59; { lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; x_17 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_14); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_12); lean_ctor_set(x_17, 2, x_16); x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 0, x_15); lean_ctor_set(x_18, 1, x_17); x_19 = lean_unsigned_to_nat(1u); x_20 = lean_nat_add(x_3, x_19); @@ -17940,29 +17940,29 @@ if (x_46 == 0) lean_object* x_47; lean_dec(x_6); x_47 = lean_box(0); -x_16 = x_41; -x_17 = x_44; +x_16 = x_38; +x_17 = x_37; x_18 = x_36; -x_19 = x_40; -x_20 = x_39; -x_21 = x_37; -x_22 = lean_box(0); -x_23 = x_42; -x_24 = x_38; +x_19 = x_44; +x_20 = lean_box(0); +x_21 = x_41; +x_22 = x_42; +x_23 = x_39; +x_24 = x_40; x_25 = x_47; goto block_28; } else { -x_16 = x_41; -x_17 = x_44; +x_16 = x_38; +x_17 = x_37; x_18 = x_36; -x_19 = x_40; -x_20 = x_39; -x_21 = x_37; -x_22 = lean_box(0); -x_23 = x_42; -x_24 = x_38; +x_19 = x_44; +x_20 = lean_box(0); +x_21 = x_41; +x_22 = x_42; +x_23 = x_39; +x_24 = x_40; x_25 = x_6; goto block_28; } @@ -18151,14 +18151,14 @@ return x_100; if (lean_obj_tag(x_18) == 0) { lean_object* x_26; -x_26 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_rintroContinue___redArg(x_1, x_2, x_3, x_25, x_17, x_7, x_4, x_8, x_21, x_24, x_20, x_19, x_16, x_23); +x_26 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_rintroContinue___redArg(x_1, x_2, x_3, x_25, x_19, x_7, x_4, x_8, x_17, x_16, x_23, x_24, x_21, x_22); return x_26; } else { lean_object* x_27; lean_dec(x_7); -x_27 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_rintroContinue___redArg(x_1, x_2, x_3, x_25, x_17, x_18, x_4, x_8, x_21, x_24, x_20, x_19, x_16, x_23); +x_27 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_rintroContinue___redArg(x_1, x_2, x_3, x_25, x_19, x_18, x_4, x_8, x_17, x_16, x_23, x_24, x_21, x_22); return x_27; } } @@ -19423,7 +19423,7 @@ goto block_154; block_59: { lean_object* x_47; -x_47 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_42, x_40, x_36, x_38, x_43); +x_47 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_39, x_41, x_45, x_35, x_42); if (lean_obj_tag(x_47) == 0) { 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; lean_object* x_55; @@ -19432,8 +19432,8 @@ lean_inc(x_48); lean_dec_ref(x_47); x_49 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_RCasesPatt_typed_x3f(x_34, x_46, x_44); lean_dec(x_44); -x_50 = l_Lean_Syntax_TSepArray_getElems___redArg(x_39); -lean_dec_ref(x_39); +x_50 = l_Lean_Syntax_TSepArray_getElems___redArg(x_40); +lean_dec_ref(x_40); x_51 = lean_array_size(x_50); x_52 = 0; x_53 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalObtain_spec__1(x_51, x_52, x_50); @@ -19442,22 +19442,22 @@ x_54 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_RCases_0__Lean_El lean_closure_set(x_54, 0, x_53); lean_closure_set(x_54, 1, x_49); lean_closure_set(x_54, 2, x_48); -x_55 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalRCases_spec__2___redArg(x_48, x_54, x_41, x_42, x_35, x_37, x_40, x_36, x_38, x_43); +x_55 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalRCases_spec__2___redArg(x_48, x_54, x_37, x_39, x_43, x_38, x_41, x_45, x_35, x_42); return x_55; } else { uint8_t x_56; lean_dec_ref(x_46); +lean_dec(x_45); lean_dec(x_44); -lean_dec(x_43); +lean_dec_ref(x_43); lean_dec(x_42); lean_dec_ref(x_41); lean_dec_ref(x_40); -lean_dec_ref(x_39); -lean_dec_ref(x_38); -lean_dec(x_37); -lean_dec(x_36); +lean_dec(x_39); +lean_dec(x_38); +lean_dec_ref(x_37); lean_dec_ref(x_35); lean_dec(x_34); x_56 = !lean_is_exclusive(x_47); @@ -19479,23 +19479,23 @@ return x_58; } block_84: { -if (lean_obj_tag(x_64) == 0) +if (lean_obj_tag(x_67) == 0) { if (lean_obj_tag(x_69) == 0) { lean_object* x_72; lean_object* x_73; lean_dec(x_70); -lean_dec_ref(x_67); -lean_dec(x_66); +lean_dec_ref(x_66); +lean_dec(x_63); lean_dec(x_62); -lean_dec_ref(x_60); +lean_dec_ref(x_61); lean_dec(x_34); x_72 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalObtain___closed__3; -x_73 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalObtain_spec__0___redArg(x_72, x_65, x_61, x_63, x_68); +x_73 = l_Lean_throwError___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalObtain_spec__0___redArg(x_72, x_64, x_68, x_60, x_65); +lean_dec(x_65); +lean_dec_ref(x_60); lean_dec(x_68); -lean_dec_ref(x_63); -lean_dec(x_61); -lean_dec_ref(x_65); +lean_dec_ref(x_64); return x_73; } else @@ -19512,14 +19512,14 @@ lean_ctor_set(x_76, 0, x_34); lean_ctor_set(x_76, 1, x_75); x_11 = x_74; x_12 = x_60; -x_13 = x_61; -x_14 = x_62; -x_15 = x_63; -x_16 = x_65; -x_17 = x_67; -x_18 = x_66; -x_19 = x_68; -x_20 = lean_box(0); +x_13 = lean_box(0); +x_14 = x_61; +x_15 = x_62; +x_16 = x_63; +x_17 = x_64; +x_18 = x_65; +x_19 = x_66; +x_20 = x_68; x_21 = x_76; goto block_29; } @@ -19535,14 +19535,14 @@ lean_inc(x_78); lean_dec_ref(x_70); x_11 = x_77; x_12 = x_60; -x_13 = x_61; -x_14 = x_62; -x_15 = x_63; -x_16 = x_65; -x_17 = x_67; -x_18 = x_66; -x_19 = x_68; -x_20 = lean_box(0); +x_13 = lean_box(0); +x_14 = x_61; +x_15 = x_62; +x_16 = x_63; +x_17 = x_64; +x_18 = x_65; +x_19 = x_66; +x_20 = x_68; x_21 = x_78; goto block_29; } @@ -19553,48 +19553,48 @@ else if (lean_obj_tag(x_70) == 0) { lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_64, 0); +x_79 = lean_ctor_get(x_67, 0); lean_inc(x_79); -lean_dec_ref(x_64); +lean_dec_ref(x_67); x_80 = l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__1; lean_inc(x_34); x_81 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_81, 0, x_34); lean_ctor_set(x_81, 1, x_80); x_35 = x_60; -x_36 = x_61; -x_37 = x_62; -x_38 = x_63; -x_39 = x_79; -x_40 = x_65; -x_41 = x_67; -x_42 = x_66; -x_43 = x_68; +x_36 = lean_box(0); +x_37 = x_61; +x_38 = x_62; +x_39 = x_63; +x_40 = x_79; +x_41 = x_64; +x_42 = x_65; +x_43 = x_66; x_44 = x_69; -x_45 = lean_box(0); +x_45 = x_68; x_46 = x_81; goto block_59; } else { lean_object* x_82; lean_object* x_83; -x_82 = lean_ctor_get(x_64, 0); +x_82 = lean_ctor_get(x_67, 0); lean_inc(x_82); -lean_dec_ref(x_64); +lean_dec_ref(x_67); x_83 = lean_ctor_get(x_70, 0); lean_inc(x_83); lean_dec_ref(x_70); x_35 = x_60; -x_36 = x_61; -x_37 = x_62; -x_38 = x_63; -x_39 = x_82; -x_40 = x_65; -x_41 = x_67; -x_42 = x_66; -x_43 = x_68; +x_36 = lean_box(0); +x_37 = x_61; +x_38 = x_62; +x_39 = x_63; +x_40 = x_82; +x_41 = x_64; +x_42 = x_65; +x_43 = x_66; x_44 = x_69; -x_45 = lean_box(0); +x_45 = x_68; x_46 = x_83; goto block_59; } @@ -19606,15 +19606,15 @@ if (lean_obj_tag(x_85) == 0) { lean_object* x_97; x_97 = lean_box(0); -x_60 = x_90; -x_61 = x_93; +x_60 = x_94; +x_61 = x_88; x_62 = x_91; -x_63 = x_94; -x_64 = x_87; -x_65 = x_92; -x_66 = x_89; -x_67 = x_88; -x_68 = x_95; +x_63 = x_89; +x_64 = x_92; +x_65 = x_95; +x_66 = x_90; +x_67 = x_87; +x_68 = x_93; x_69 = x_86; x_70 = x_97; x_71 = lean_box(0); @@ -19636,15 +19636,15 @@ x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); lean_dec_ref(x_100); lean_ctor_set(x_85, 0, x_101); -x_60 = x_90; -x_61 = x_93; +x_60 = x_94; +x_61 = x_88; x_62 = x_91; -x_63 = x_94; -x_64 = x_87; -x_65 = x_92; -x_66 = x_89; -x_67 = x_88; -x_68 = x_95; +x_63 = x_89; +x_64 = x_92; +x_65 = x_95; +x_66 = x_90; +x_67 = x_87; +x_68 = x_93; x_69 = x_86; x_70 = x_85; x_71 = lean_box(0); @@ -19697,15 +19697,15 @@ lean_inc(x_107); lean_dec_ref(x_106); x_108 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_108, 0, x_107); -x_60 = x_90; -x_61 = x_93; +x_60 = x_94; +x_61 = x_88; x_62 = x_91; -x_63 = x_94; -x_64 = x_87; -x_65 = x_92; -x_66 = x_89; -x_67 = x_88; -x_68 = x_95; +x_63 = x_89; +x_64 = x_92; +x_65 = x_95; +x_66 = x_90; +x_67 = x_87; +x_68 = x_93; x_69 = x_86; x_70 = x_108; x_71 = lean_box(0); @@ -19756,7 +19756,7 @@ if (x_128 == 0) { uint8_t x_129; lean_inc(x_127); -x_129 = l_Lean_Syntax_matchesNull(x_127, x_115); +x_129 = l_Lean_Syntax_matchesNull(x_127, x_114); if (x_129 == 0) { lean_object* x_130; @@ -19770,7 +19770,7 @@ lean_dec_ref(x_119); lean_dec(x_118); lean_dec_ref(x_117); lean_dec(x_116); -lean_dec(x_114); +lean_dec(x_115); lean_dec(x_34); x_130 = l_Lean_Elab_throwUnsupportedSyntax___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalRCases_spec__0___redArg(); return x_130; @@ -19784,7 +19784,7 @@ x_132 = l_Lean_Syntax_getArgs(x_131); lean_dec(x_131); x_133 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_133, 0, x_132); -x_85 = x_114; +x_85 = x_115; x_86 = x_116; x_87 = x_133; x_88 = x_117; @@ -19804,7 +19804,7 @@ else lean_object* x_134; lean_dec(x_127); x_134 = lean_box(0); -x_85 = x_114; +x_85 = x_115; x_86 = x_116; x_87 = x_134; x_88 = x_117; @@ -19855,8 +19855,8 @@ x_151 = l_Lean_Syntax_getArg(x_147, x_113); lean_dec(x_147); x_152 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_152, 0, x_151); -x_114 = x_136; -x_115 = x_146; +x_114 = x_146; +x_115 = x_136; x_116 = x_152; x_117 = x_137; x_118 = x_138; @@ -19875,8 +19875,8 @@ else lean_object* x_153; lean_dec(x_147); x_153 = lean_box(0); -x_114 = x_136; -x_115 = x_146; +x_114 = x_146; +x_115 = x_136; x_116 = x_153; x_117 = x_137; x_118 = x_138; @@ -19894,7 +19894,7 @@ goto block_135; block_29: { lean_object* x_22; -x_22 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_18, x_16, x_13, x_15, x_19); +x_22 = l_Lean_Elab_Tactic_getMainGoal___redArg(x_16, x_17, x_20, x_12, x_18); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -19906,20 +19906,20 @@ x_24 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_RCases_0__Lean_El lean_closure_set(x_24, 0, x_21); lean_closure_set(x_24, 1, x_11); lean_closure_set(x_24, 2, x_23); -x_25 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalRCases_spec__2___redArg(x_23, x_24, x_17, x_18, x_12, x_14, x_16, x_13, x_15, x_19); +x_25 = l_Lean_MVarId_withContext___at___00__private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_evalRCases_spec__2___redArg(x_23, x_24, x_14, x_16, x_19, x_15, x_17, x_20, x_12, x_18); return x_25; } else { uint8_t x_26; lean_dec_ref(x_21); -lean_dec(x_19); +lean_dec(x_20); +lean_dec_ref(x_19); lean_dec(x_18); lean_dec_ref(x_17); -lean_dec_ref(x_16); -lean_dec_ref(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_dec(x_16); +lean_dec(x_15); +lean_dec_ref(x_14); lean_dec_ref(x_12); lean_dec(x_11); x_26 = !lean_is_exclusive(x_22); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/AC/Eq.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/AC/Eq.c index 0f7ecce489f1..8c06c5b00662 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/AC/Eq.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/AC/Eq.c @@ -20480,9 +20480,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 { @@ -20501,7 +20501,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; } @@ -20527,8 +20527,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; } @@ -20536,8 +20536,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; } @@ -20954,9 +20954,9 @@ goto block_154; block_147: { lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_nat_add(x_142, x_143); +x_144 = lean_nat_add(x_141, x_143); lean_dec(x_143); -lean_dec(x_142); +lean_dec(x_141); if (lean_is_scalar(x_138)) { x_145 = lean_alloc_ctor(0, 5, 0); } else { @@ -20975,7 +20975,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_141); +lean_ctor_set(x_146, 3, x_142); lean_ctor_set(x_146, 4, x_145); return x_146; } @@ -21001,8 +21001,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_150; -x_142 = x_151; +x_141 = x_151; +x_142 = x_150; x_143 = x_152; goto block_147; } @@ -21010,8 +21010,8 @@ else { lean_object* x_153; x_153 = lean_unsigned_to_nat(0u); -x_141 = x_150; -x_142 = x_151; +x_141 = x_151; +x_142 = x_150; x_143 = x_153; goto block_147; } @@ -22175,7 +22175,7 @@ return x_33; { lean_object* x_45; lean_object* x_46; x_45 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__3_spec__3___redArg___closed__3; -x_46 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__0(x_45, x_41, x_40, x_43, x_38, x_39, x_42, x_35, x_37, x_36); +x_46 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__0(x_45, x_41, x_38, x_42, x_39, x_40, x_44, x_37, x_43, x_35); x_20 = x_46; goto block_34; } @@ -22183,17 +22183,17 @@ goto block_34; { lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; x_67 = lean_ctor_get(x_3, 1); -lean_inc_ref(x_54); +lean_inc_ref(x_55); lean_inc_ref(x_67); -x_68 = l_Lean_Grind_AC_Seq_union(x_67, x_54); -lean_inc_ref(x_56); +x_68 = l_Lean_Grind_AC_Seq_union(x_67, x_55); +lean_inc_ref(x_54); lean_inc_ref(x_53); -x_69 = l_Lean_Grind_AC_Seq_union(x_53, x_56); +x_69 = l_Lean_Grind_AC_Seq_union(x_53, x_54); lean_inc_ref(x_3); x_70 = lean_alloc_ctor(9, 5, 0); -lean_ctor_set(x_70, 0, x_56); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_54); +lean_ctor_set(x_70, 0, x_54); +lean_ctor_set(x_70, 1, x_56); +lean_ctor_set(x_70, 2, x_55); lean_ctor_set(x_70, 3, x_3); lean_ctor_set(x_70, 4, x_17); lean_inc(x_57); @@ -22315,16 +22315,16 @@ if (x_91 == 0) if (x_89 == 0) { lean_dec(x_17); -x_35 = x_85; -x_36 = x_87; -x_37 = x_86; -x_38 = x_82; -x_39 = x_83; -x_40 = x_80; +x_35 = x_87; +x_36 = lean_box(0); +x_37 = x_85; +x_38 = x_80; +x_39 = x_82; +x_40 = x_83; x_41 = x_79; -x_42 = x_84; -x_43 = x_81; -x_44 = lean_box(0); +x_42 = x_81; +x_43 = x_86; +x_44 = x_84; goto block_47; } else @@ -22388,9 +22388,9 @@ x_102 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Le x_103 = l_Lean_Option_get___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__2(x_101, x_102); if (x_103 == 0) { -x_54 = x_100; -x_55 = x_99; -x_56 = x_98; +x_54 = x_98; +x_55 = x_100; +x_56 = x_99; x_57 = x_79; x_58 = x_80; x_59 = x_81; @@ -22445,9 +22445,9 @@ goto block_34; } else { -x_54 = x_100; -x_55 = x_99; -x_56 = x_98; +x_54 = x_98; +x_55 = x_100; +x_56 = x_99; x_57 = x_79; x_58 = x_80; x_59 = x_81; @@ -22468,16 +22468,16 @@ goto block_78; else { lean_dec(x_17); -x_35 = x_85; -x_36 = x_87; -x_37 = x_86; -x_38 = x_82; -x_39 = x_83; -x_40 = x_80; +x_35 = x_87; +x_36 = lean_box(0); +x_37 = x_85; +x_38 = x_80; +x_39 = x_82; +x_40 = x_83; x_41 = x_79; -x_42 = x_84; -x_43 = x_81; -x_44 = lean_box(0); +x_42 = x_81; +x_43 = x_86; +x_44 = x_84; goto block_47; } } @@ -22849,7 +22849,7 @@ return x_34; { lean_object* x_46; lean_object* x_47; x_46 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__3_spec__3___redArg___closed__3; -x_47 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__0(x_46, x_37, x_36, x_45, x_44, x_40, x_38, x_39, x_43, x_42); +x_47 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__0(x_46, x_41, x_38, x_37, x_40, x_43, x_39, x_44, x_36, x_42); x_21 = x_47; goto block_35; } @@ -22981,16 +22981,16 @@ if (x_92 == 0) if (x_90 == 0) { lean_dec(x_18); -x_36 = x_81; -x_37 = x_80; -x_38 = x_85; -x_39 = x_86; -x_40 = x_84; -x_41 = lean_box(0); +x_36 = x_87; +x_37 = x_82; +x_38 = x_81; +x_39 = x_85; +x_40 = x_83; +x_41 = x_80; x_42 = x_88; -x_43 = x_87; -x_44 = x_83; -x_45 = x_82; +x_43 = x_84; +x_44 = x_86; +x_45 = lean_box(0); goto block_48; } else @@ -23130,16 +23130,16 @@ goto block_79; else { lean_dec(x_18); -x_36 = x_81; -x_37 = x_80; -x_38 = x_85; -x_39 = x_86; -x_40 = x_84; -x_41 = lean_box(0); +x_36 = x_87; +x_37 = x_82; +x_38 = x_81; +x_39 = x_85; +x_40 = x_83; +x_41 = x_80; x_42 = x_88; -x_43 = x_87; -x_44 = x_83; -x_45 = x_82; +x_43 = x_84; +x_44 = x_86; +x_45 = lean_box(0); goto block_48; } } @@ -27846,9 +27846,9 @@ goto block_49; block_42: { lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_nat_add(x_36, x_38); +x_39 = lean_nat_add(x_37, x_38); lean_dec(x_38); -lean_dec(x_36); +lean_dec(x_37); if (lean_is_scalar(x_33)) { x_40 = lean_alloc_ctor(0, 5, 0); } else { @@ -27867,7 +27867,7 @@ if (lean_is_scalar(x_23)) { lean_ctor_set(x_41, 0, x_35); lean_ctor_set(x_41, 1, x_25); lean_ctor_set(x_41, 2, x_26); -lean_ctor_set(x_41, 3, x_37); +lean_ctor_set(x_41, 3, x_36); lean_ctor_set(x_41, 4, x_40); return x_41; } @@ -27893,8 +27893,8 @@ if (lean_obj_tag(x_28) == 0) lean_object* x_47; x_47 = lean_ctor_get(x_28, 0); lean_inc(x_47); -x_36 = x_46; -x_37 = x_45; +x_36 = x_45; +x_37 = x_46; x_38 = x_47; goto block_42; } @@ -27902,8 +27902,8 @@ else { lean_object* x_48; x_48 = lean_unsigned_to_nat(0u); -x_36 = x_46; -x_37 = x_45; +x_36 = x_45; +x_37 = x_46; x_38 = x_48; goto block_42; } @@ -28477,9 +28477,9 @@ goto block_188; block_181: { lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_178 = lean_nat_add(x_176, x_177); +x_178 = lean_nat_add(x_175, x_177); lean_dec(x_177); -lean_dec(x_176); +lean_dec(x_175); if (lean_is_scalar(x_172)) { x_179 = lean_alloc_ctor(0, 5, 0); } else { @@ -28498,7 +28498,7 @@ if (lean_is_scalar(x_162)) { lean_ctor_set(x_180, 0, x_174); lean_ctor_set(x_180, 1, x_164); lean_ctor_set(x_180, 2, x_165); -lean_ctor_set(x_180, 3, x_175); +lean_ctor_set(x_180, 3, x_176); lean_ctor_set(x_180, 4, x_179); return x_180; } @@ -28524,8 +28524,8 @@ if (lean_obj_tag(x_167) == 0) lean_object* x_186; x_186 = lean_ctor_get(x_167, 0); lean_inc(x_186); -x_175 = x_184; -x_176 = x_185; +x_175 = x_185; +x_176 = x_184; x_177 = x_186; goto block_181; } @@ -28533,8 +28533,8 @@ else { lean_object* x_187; x_187 = lean_unsigned_to_nat(0u); -x_175 = x_184; -x_176 = x_185; +x_175 = x_185; +x_176 = x_184; x_177 = x_187; goto block_181; } @@ -36437,7 +36437,7 @@ x_110 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_110, 0, x_109); x_111 = l_Lean_MessageData_ofFormat(x_110); x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_98); +lean_ctor_set(x_112, 0, x_104); lean_ctor_set(x_112, 1, x_111); x_113 = l_List_forIn_x27_loop___at___00List_forIn_x27_loop___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_EqCnstr_superposeWithAC_spec__3_spec__3___redArg___closed__16; x_114 = lean_alloc_ctor(7, 2, 0); @@ -36456,24 +36456,24 @@ x_119 = l_Lean_MessageData_ofExpr(x_95); x_120 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_120, 0, x_118); lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_propagateEqs_process_spec__9___redArg(x_16, x_120, x_101, x_104, x_97, x_100, x_103); +x_121 = l_Lean_addTrace___at___00__private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC_propagateEqs_process_spec__9___redArg(x_16, x_120, x_97, x_96, x_99, x_98, x_100); x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); lean_dec_ref(x_121); x_123 = lean_ctor_get(x_122, 1); lean_inc(x_123); lean_dec(x_122); -x_26 = x_106; +x_26 = x_107; x_27 = x_123; -x_28 = x_107; -x_29 = x_99; -x_30 = x_105; -x_31 = x_96; -x_32 = x_108; -x_33 = x_104; -x_34 = x_97; -x_35 = x_100; -x_36 = x_103; +x_28 = x_105; +x_29 = x_103; +x_30 = x_106; +x_31 = x_108; +x_32 = x_102; +x_33 = x_96; +x_34 = x_99; +x_35 = x_98; +x_36 = x_100; x_37 = lean_box(0); goto block_94; } @@ -36602,19 +36602,19 @@ if (x_163 == 0) lean_object* x_164; x_164 = l___private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC___aux__Lean__Meta__Tactic__Grind__AC__Eq______macroRules____private__Lean__Meta__Tactic__Grind__AC__Eq__0__Lean__Meta__Grind__AC__commandGen__cnstr__fns____1___closed__253; x_95 = x_159; -x_96 = x_130; -x_97 = x_133; -x_98 = x_162; -x_99 = x_128; -x_100 = x_134; -x_101 = x_158; -x_102 = lean_box(0); -x_103 = x_135; -x_104 = x_132; -x_105 = x_129; -x_106 = x_152; -x_107 = x_127; -x_108 = x_131; +x_96 = x_132; +x_97 = x_158; +x_98 = x_134; +x_99 = x_133; +x_100 = x_135; +x_101 = lean_box(0); +x_102 = x_131; +x_103 = x_128; +x_104 = x_162; +x_105 = x_127; +x_106 = x_129; +x_107 = x_152; +x_108 = x_130; x_109 = x_164; goto block_124; } @@ -36623,19 +36623,19 @@ else lean_object* x_165; x_165 = l___private_Lean_Meta_Tactic_Grind_AC_Eq_0__Lean_Meta_Grind_AC___aux__Lean__Meta__Tactic__Grind__AC__Eq______macroRules____private__Lean__Meta__Tactic__Grind__AC__Eq__0__Lean__Meta__Grind__AC__commandGen__cnstr__fns____1___closed__220; x_95 = x_159; -x_96 = x_130; -x_97 = x_133; -x_98 = x_162; -x_99 = x_128; -x_100 = x_134; -x_101 = x_158; -x_102 = lean_box(0); -x_103 = x_135; -x_104 = x_132; -x_105 = x_129; -x_106 = x_152; -x_107 = x_127; -x_108 = x_131; +x_96 = x_132; +x_97 = x_158; +x_98 = x_134; +x_99 = x_133; +x_100 = x_135; +x_101 = lean_box(0); +x_102 = x_131; +x_103 = x_128; +x_104 = x_162; +x_105 = x_127; +x_106 = x_129; +x_107 = x_152; +x_108 = x_130; x_109 = x_165; goto block_124; } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Arith/Cutsat/Proof.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Arith/Cutsat/Proof.c index 9161abf187b5..a43f00c63627 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Arith/Cutsat/Proof.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Arith/Cutsat/Proof.c @@ -10652,7 +10652,7 @@ goto block_77; { uint8_t x_18; lean_object* x_19; lean_object* x_20; x_18 = 0; -x_19 = l_Lean_Expr_letE___override(x_17, x_14, x_13, x_15, x_18); +x_19 = l_Lean_Expr_letE___override(x_17, x_13, x_14, x_15, x_18); x_20 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_20, 0, x_19); return x_20; @@ -10661,11 +10661,11 @@ return x_20; { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; x_42 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__6; -x_43 = l_Lean_Name_mkStr3(x_39, x_37, x_42); -lean_inc(x_38); -x_44 = l_Lean_mkConst(x_43, x_38); -x_45 = l_Lean_Meta_Grind_mkLetOfMap___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext_spec__17(x_5, x_40, x_41, x_44, x_33); -lean_dec_ref(x_40); +x_43 = l_Lean_Name_mkStr3(x_40, x_38, x_42); +lean_inc(x_37); +x_44 = l_Lean_mkConst(x_43, x_37); +x_45 = l_Lean_Meta_Grind_mkLetOfMap___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext_spec__17(x_5, x_39, x_41, x_44, x_33); +lean_dec_ref(x_39); x_46 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__7; x_47 = lean_array_push(x_46, x_1); x_48 = lean_expr_abstract(x_45, x_47); @@ -10675,7 +10675,7 @@ x_49 = l_Lean_Expr_hasLooseBVars(x_48); if (x_49 == 0) { lean_object* x_50; -lean_dec(x_38); +lean_dec(x_37); lean_dec_ref(x_36); lean_dec(x_11); lean_dec_ref(x_10); @@ -10700,15 +10700,15 @@ x_54 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_ x_55 = lean_box(0); x_56 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_38); +lean_ctor_set(x_56, 1, x_37); x_57 = l_Lean_mkConst(x_54, x_56); x_58 = l_Lean_Expr_app___override(x_57, x_51); if (x_2 == 0) { lean_object* x_59; x_59 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__12; -x_13 = x_53; -x_14 = x_58; +x_13 = x_58; +x_14 = x_53; x_15 = x_48; x_16 = lean_box(0); x_17 = x_59; @@ -10718,8 +10718,8 @@ else { lean_object* x_60; x_60 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__14; -x_13 = x_53; -x_14 = x_58; +x_13 = x_58; +x_14 = x_53; x_15 = x_48; x_16 = lean_box(0); x_17 = x_60; @@ -10729,7 +10729,7 @@ goto block_21; else { lean_dec_ref(x_48); -lean_dec(x_38); +lean_dec(x_37); return x_52; } } @@ -10747,10 +10747,10 @@ if (x_2 == 0) { lean_object* x_75; x_75 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__21; -x_37 = x_71; -x_38 = x_72; -x_39 = x_70; -x_40 = x_74; +x_37 = x_72; +x_38 = x_71; +x_39 = x_74; +x_40 = x_70; x_41 = x_75; goto block_61; } @@ -10758,10 +10758,10 @@ else { lean_object* x_76; x_76 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__23; -x_37 = x_71; -x_38 = x_72; -x_39 = x_70; -x_40 = x_74; +x_37 = x_72; +x_38 = x_71; +x_39 = x_74; +x_40 = x_70; x_41 = x_76; goto block_61; } @@ -14889,9 +14889,9 @@ goto block_380; { lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_55 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_56 = l_Lean_mkApp8(x_50, x_46, x_45, x_49, x_48, x_54, x_52, x_51, x_55); +x_56 = l_Lean_mkApp8(x_49, x_46, x_45, x_48, x_53, x_54, x_47, x_52, x_55); x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_53); +lean_ctor_set(x_57, 0, x_50); lean_ctor_set(x_57, 1, x_56); if (lean_is_scalar(x_44)) { x_58 = lean_alloc_ctor(0, 1, 0); @@ -14905,53 +14905,53 @@ return x_58; { lean_object* x_69; uint8_t x_70; x_69 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_70 = lean_int_dec_le(x_69, x_67); +x_70 = lean_int_dec_le(x_69, x_63); if (x_70 == 0) { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; x_71 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_72 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_60); +lean_inc(x_64); x_73 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_60); +lean_ctor_set(x_73, 1, x_64); x_74 = l_Lean_Expr_const___override(x_71, x_73); -lean_inc_ref(x_62); -x_75 = l_Lean_Name_mkStr1(x_62); -lean_inc(x_60); -x_76 = l_Lean_Expr_const___override(x_75, x_60); +lean_inc_ref(x_66); +x_75 = l_Lean_Name_mkStr1(x_66); +lean_inc(x_64); +x_76 = l_Lean_Expr_const___override(x_75, x_64); x_77 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_78 = l_Lean_Name_mkStr2(x_62, x_77); -x_79 = l_Lean_Expr_const___override(x_78, x_60); -x_80 = lean_int_neg(x_67); +x_78 = l_Lean_Name_mkStr2(x_66, x_77); +x_79 = l_Lean_Expr_const___override(x_78, x_64); +x_80 = lean_int_neg(x_63); x_81 = l_Int_toNat(x_80); lean_dec(x_80); x_82 = l_Lean_instToExprInt_mkNat(x_81); x_83 = l_Lean_mkApp3(x_74, x_76, x_79, x_82); -x_47 = lean_box(0); -x_48 = x_68; -x_49 = x_63; -x_50 = x_64; -x_51 = x_66; -x_52 = x_65; -x_53 = x_67; +x_47 = x_60; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = lean_box(0); +x_52 = x_67; +x_53 = x_68; x_54 = x_83; goto block_59; } else { lean_object* x_84; lean_object* x_85; -lean_dec_ref(x_62); -lean_dec(x_60); -x_84 = l_Int_toNat(x_67); +lean_dec_ref(x_66); +lean_dec(x_64); +x_84 = l_Int_toNat(x_63); x_85 = l_Lean_instToExprInt_mkNat(x_84); -x_47 = lean_box(0); -x_48 = x_68; -x_49 = x_63; -x_50 = x_64; -x_51 = x_66; -x_52 = x_65; -x_53 = x_67; +x_47 = x_60; +x_48 = x_61; +x_49 = x_62; +x_50 = x_63; +x_51 = lean_box(0); +x_52 = x_67; +x_53 = x_68; x_54 = x_85; goto block_59; } @@ -14960,57 +14960,57 @@ goto block_59; { lean_object* x_96; uint8_t x_97; x_96 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_97 = lean_int_dec_le(x_96, x_88); +x_97 = lean_int_dec_le(x_96, x_91); if (x_97 == 0) { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; x_98 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_99 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_87); +lean_inc(x_90); x_100 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_87); +lean_ctor_set(x_100, 1, x_90); x_101 = l_Lean_Expr_const___override(x_98, x_100); -lean_inc_ref(x_90); -x_102 = l_Lean_Name_mkStr1(x_90); -lean_inc(x_87); -x_103 = l_Lean_Expr_const___override(x_102, x_87); +lean_inc_ref(x_94); +x_102 = l_Lean_Name_mkStr1(x_94); +lean_inc(x_90); +x_103 = l_Lean_Expr_const___override(x_102, x_90); x_104 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_90); -x_105 = l_Lean_Name_mkStr2(x_90, x_104); -lean_inc(x_87); -x_106 = l_Lean_Expr_const___override(x_105, x_87); -x_107 = lean_int_neg(x_88); -lean_dec(x_88); +lean_inc_ref(x_94); +x_105 = l_Lean_Name_mkStr2(x_94, x_104); +lean_inc(x_90); +x_106 = l_Lean_Expr_const___override(x_105, x_90); +x_107 = lean_int_neg(x_91); +lean_dec(x_91); x_108 = l_Int_toNat(x_107); lean_dec(x_107); x_109 = l_Lean_instToExprInt_mkNat(x_108); x_110 = l_Lean_mkApp3(x_101, x_103, x_106, x_109); x_60 = x_87; -x_61 = lean_box(0); -x_62 = x_90; -x_63 = x_95; -x_64 = x_91; -x_65 = x_93; -x_66 = x_92; -x_67 = x_94; +x_61 = x_95; +x_62 = x_88; +x_63 = x_89; +x_64 = x_90; +x_65 = lean_box(0); +x_66 = x_94; +x_67 = x_93; x_68 = x_110; goto block_86; } else { lean_object* x_111; lean_object* x_112; -x_111 = l_Int_toNat(x_88); -lean_dec(x_88); +x_111 = l_Int_toNat(x_91); +lean_dec(x_91); x_112 = l_Lean_instToExprInt_mkNat(x_111); x_60 = x_87; -x_61 = lean_box(0); -x_62 = x_90; -x_63 = x_95; -x_64 = x_91; -x_65 = x_93; -x_66 = x_92; -x_67 = x_94; +x_61 = x_95; +x_62 = x_88; +x_63 = x_89; +x_64 = x_90; +x_65 = lean_box(0); +x_66 = x_94; +x_67 = x_93; x_68 = x_112; goto block_86; } @@ -15019,11 +15019,11 @@ goto block_86; { lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; x_123 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -lean_inc_ref(x_120); -x_124 = l_Lean_mkApp9(x_116, x_46, x_45, x_121, x_115, x_120, x_122, x_118, x_119, x_123); +lean_inc_ref(x_115); +x_124 = l_Lean_mkApp9(x_116, x_46, x_45, x_121, x_120, x_115, x_122, x_114, x_117, x_123); x_125 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_125, 0, x_114); -lean_ctor_set(x_125, 1, x_120); +lean_ctor_set(x_125, 0, x_119); +lean_ctor_set(x_125, 1, x_115); lean_ctor_set(x_125, 2, x_124); x_126 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_126, 0, x_125); @@ -15033,36 +15033,36 @@ return x_126; { lean_object* x_138; uint8_t x_139; x_138 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_139 = lean_int_dec_le(x_138, x_128); +x_139 = lean_int_dec_le(x_138, x_133); 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; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; x_140 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_141 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_133); +lean_inc(x_135); x_142 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_133); +lean_ctor_set(x_142, 1, x_135); x_143 = l_Lean_Expr_const___override(x_140, x_142); -lean_inc_ref(x_131); -x_144 = l_Lean_Name_mkStr1(x_131); -lean_inc(x_133); -x_145 = l_Lean_Expr_const___override(x_144, x_133); +lean_inc_ref(x_128); +x_144 = l_Lean_Name_mkStr1(x_128); +lean_inc(x_135); +x_145 = l_Lean_Expr_const___override(x_144, x_135); x_146 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_147 = l_Lean_Name_mkStr2(x_131, x_146); -x_148 = l_Lean_Expr_const___override(x_147, x_133); -x_149 = lean_int_neg(x_128); +x_147 = l_Lean_Name_mkStr2(x_128, x_146); +x_148 = l_Lean_Expr_const___override(x_147, x_135); +x_149 = lean_int_neg(x_133); x_150 = l_Int_toNat(x_149); lean_dec(x_149); x_151 = l_Lean_instToExprInt_mkNat(x_150); x_152 = l_Lean_mkApp3(x_143, x_145, x_148, x_151); -x_114 = x_128; -x_115 = x_137; -x_116 = x_129; -x_117 = lean_box(0); -x_118 = x_132; -x_119 = x_135; -x_120 = x_134; +x_114 = x_129; +x_115 = x_131; +x_116 = x_130; +x_117 = x_132; +x_118 = lean_box(0); +x_119 = x_133; +x_120 = x_137; x_121 = x_136; x_122 = x_152; goto block_127; @@ -15070,17 +15070,17 @@ goto block_127; else { lean_object* x_153; lean_object* x_154; -lean_dec(x_133); -lean_dec_ref(x_131); -x_153 = l_Int_toNat(x_128); +lean_dec(x_135); +lean_dec_ref(x_128); +x_153 = l_Int_toNat(x_133); x_154 = l_Lean_instToExprInt_mkNat(x_153); -x_114 = x_128; -x_115 = x_137; -x_116 = x_129; -x_117 = lean_box(0); -x_118 = x_132; -x_119 = x_135; -x_120 = x_134; +x_114 = x_129; +x_115 = x_131; +x_116 = x_130; +x_117 = x_132; +x_118 = lean_box(0); +x_119 = x_133; +x_120 = x_137; x_121 = x_136; x_122 = x_154; goto block_127; @@ -15090,7 +15090,7 @@ goto block_127; { lean_object* x_166; uint8_t x_167; x_166 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_167 = lean_int_dec_le(x_166, x_161); +x_167 = lean_int_dec_le(x_166, x_163); 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; lean_object* x_178; lean_object* x_179; lean_object* x_180; @@ -15101,29 +15101,29 @@ x_170 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_170, 0, x_169); lean_ctor_set(x_170, 1, x_164); x_171 = l_Lean_Expr_const___override(x_168, x_170); -lean_inc_ref(x_159); -x_172 = l_Lean_Name_mkStr1(x_159); +lean_inc_ref(x_156); +x_172 = l_Lean_Name_mkStr1(x_156); lean_inc(x_164); x_173 = l_Lean_Expr_const___override(x_172, x_164); x_174 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_159); -x_175 = l_Lean_Name_mkStr2(x_159, x_174); +lean_inc_ref(x_156); +x_175 = l_Lean_Name_mkStr2(x_156, x_174); lean_inc(x_164); x_176 = l_Lean_Expr_const___override(x_175, x_164); -x_177 = lean_int_neg(x_161); -lean_dec(x_161); +x_177 = lean_int_neg(x_163); +lean_dec(x_163); x_178 = l_Int_toNat(x_177); lean_dec(x_177); x_179 = l_Lean_instToExprInt_mkNat(x_178); x_180 = l_Lean_mkApp3(x_171, x_173, x_176, x_179); x_128 = x_156; x_129 = x_157; -x_130 = lean_box(0); -x_131 = x_159; +x_130 = x_159; +x_131 = x_158; x_132 = x_160; -x_133 = x_164; -x_134 = x_163; -x_135 = x_162; +x_133 = x_162; +x_134 = lean_box(0); +x_135 = x_164; x_136 = x_165; x_137 = x_180; goto block_155; @@ -15131,17 +15131,17 @@ goto block_155; else { lean_object* x_181; lean_object* x_182; -x_181 = l_Int_toNat(x_161); -lean_dec(x_161); +x_181 = l_Int_toNat(x_163); +lean_dec(x_163); x_182 = l_Lean_instToExprInt_mkNat(x_181); x_128 = x_156; x_129 = x_157; -x_130 = lean_box(0); -x_131 = x_159; +x_130 = x_159; +x_131 = x_158; x_132 = x_160; -x_133 = x_164; -x_134 = x_163; -x_135 = x_162; +x_133 = x_162; +x_134 = lean_box(0); +x_135 = x_164; x_136 = x_165; x_137 = x_182; goto block_155; @@ -15151,11 +15151,11 @@ goto block_155; { lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; x_193 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -lean_inc_ref(x_189); -x_194 = l_Lean_mkApp9(x_187, x_46, x_45, x_186, x_189, x_191, x_192, x_188, x_185, x_193); +lean_inc_ref(x_190); +x_194 = l_Lean_mkApp9(x_184, x_46, x_45, x_187, x_190, x_191, x_192, x_186, x_189, x_193); x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_190); -lean_ctor_set(x_195, 1, x_189); +lean_ctor_set(x_195, 0, x_185); +lean_ctor_set(x_195, 1, x_190); lean_ctor_set(x_195, 2, x_194); x_196 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_196, 0, x_195); @@ -15165,36 +15165,36 @@ return x_196; { lean_object* x_208; uint8_t x_209; x_208 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_209 = lean_int_dec_le(x_208, x_205); +x_209 = lean_int_dec_le(x_208, x_201); if (x_209 == 0) { lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; x_210 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_211 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_206); +lean_inc(x_202); x_212 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_212, 0, x_211); -lean_ctor_set(x_212, 1, x_206); +lean_ctor_set(x_212, 1, x_202); x_213 = l_Lean_Expr_const___override(x_210, x_212); -lean_inc_ref(x_204); -x_214 = l_Lean_Name_mkStr1(x_204); -lean_inc(x_206); -x_215 = l_Lean_Expr_const___override(x_214, x_206); +lean_inc_ref(x_199); +x_214 = l_Lean_Name_mkStr1(x_199); +lean_inc(x_202); +x_215 = l_Lean_Expr_const___override(x_214, x_202); x_216 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_217 = l_Lean_Name_mkStr2(x_204, x_216); -x_218 = l_Lean_Expr_const___override(x_217, x_206); -x_219 = lean_int_neg(x_205); +x_217 = l_Lean_Name_mkStr2(x_199, x_216); +x_218 = l_Lean_Expr_const___override(x_217, x_202); +x_219 = lean_int_neg(x_201); x_220 = l_Int_toNat(x_219); lean_dec(x_219); x_221 = l_Lean_instToExprInt_mkNat(x_220); x_222 = l_Lean_mkApp3(x_213, x_215, x_218, x_221); -x_184 = lean_box(0); -x_185 = x_200; -x_186 = x_199; -x_187 = x_201; -x_188 = x_203; -x_189 = x_202; -x_190 = x_205; +x_184 = x_198; +x_185 = x_201; +x_186 = x_200; +x_187 = x_203; +x_188 = lean_box(0); +x_189 = x_204; +x_190 = x_206; x_191 = x_207; x_192 = x_222; goto block_197; @@ -15202,17 +15202,17 @@ goto block_197; else { lean_object* x_223; lean_object* x_224; -lean_dec(x_206); -lean_dec_ref(x_204); -x_223 = l_Int_toNat(x_205); +lean_dec(x_202); +lean_dec_ref(x_199); +x_223 = l_Int_toNat(x_201); x_224 = l_Lean_instToExprInt_mkNat(x_223); -x_184 = lean_box(0); -x_185 = x_200; -x_186 = x_199; -x_187 = x_201; -x_188 = x_203; -x_189 = x_202; -x_190 = x_205; +x_184 = x_198; +x_185 = x_201; +x_186 = x_200; +x_187 = x_203; +x_188 = lean_box(0); +x_189 = x_204; +x_190 = x_206; x_191 = x_207; x_192 = x_224; goto block_197; @@ -15222,40 +15222,40 @@ goto block_197; { lean_object* x_236; uint8_t x_237; x_236 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_237 = lean_int_dec_le(x_236, x_233); +x_237 = lean_int_dec_le(x_236, x_231); 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 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_239 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_234); +lean_inc(x_230); x_240 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_240, 0, x_239); -lean_ctor_set(x_240, 1, x_234); +lean_ctor_set(x_240, 1, x_230); x_241 = l_Lean_Expr_const___override(x_238, x_240); -lean_inc_ref(x_232); -x_242 = l_Lean_Name_mkStr1(x_232); -lean_inc(x_234); -x_243 = l_Lean_Expr_const___override(x_242, x_234); +lean_inc_ref(x_229); +x_242 = l_Lean_Name_mkStr1(x_229); +lean_inc(x_230); +x_243 = l_Lean_Expr_const___override(x_242, x_230); x_244 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_232); -x_245 = l_Lean_Name_mkStr2(x_232, x_244); -lean_inc(x_234); -x_246 = l_Lean_Expr_const___override(x_245, x_234); -x_247 = lean_int_neg(x_233); -lean_dec(x_233); +lean_inc_ref(x_229); +x_245 = l_Lean_Name_mkStr2(x_229, x_244); +lean_inc(x_230); +x_246 = l_Lean_Expr_const___override(x_245, x_230); +x_247 = lean_int_neg(x_231); +lean_dec(x_231); x_248 = l_Int_toNat(x_247); lean_dec(x_247); x_249 = l_Lean_instToExprInt_mkNat(x_248); x_250 = l_Lean_mkApp3(x_241, x_243, x_246, x_249); -x_198 = lean_box(0); -x_199 = x_235; -x_200 = x_227; -x_201 = x_228; +x_198 = x_226; +x_199 = x_229; +x_200 = x_228; +x_201 = x_227; x_202 = x_230; -x_203 = x_229; -x_204 = x_232; -x_205 = x_231; +x_203 = x_235; +x_204 = x_233; +x_205 = lean_box(0); x_206 = x_234; x_207 = x_250; goto block_225; @@ -15263,17 +15263,17 @@ goto block_225; else { lean_object* x_251; lean_object* x_252; -x_251 = l_Int_toNat(x_233); -lean_dec(x_233); +x_251 = l_Int_toNat(x_231); +lean_dec(x_231); x_252 = l_Lean_instToExprInt_mkNat(x_251); -x_198 = lean_box(0); -x_199 = x_235; -x_200 = x_227; -x_201 = x_228; +x_198 = x_226; +x_199 = x_229; +x_200 = x_228; +x_201 = x_227; x_202 = x_230; -x_203 = x_229; -x_204 = x_232; -x_205 = x_231; +x_203 = x_235; +x_204 = x_233; +x_205 = lean_box(0); x_206 = x_234; x_207 = x_252; goto block_225; @@ -15281,23 +15281,23 @@ goto block_225; } block_316: { -switch (lean_obj_tag(x_255)) { +switch (lean_obj_tag(x_254)) { case 0: { -switch (lean_obj_tag(x_254)) { +switch (lean_obj_tag(x_255)) { case 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; uint8_t x_266; -x_257 = lean_ctor_get(x_255, 0); +x_257 = lean_ctor_get(x_254, 0); lean_inc(x_257); -x_258 = lean_ctor_get(x_255, 1); +x_258 = lean_ctor_get(x_254, 1); lean_inc_ref(x_258); -lean_dec_ref(x_255); -x_259 = lean_ctor_get(x_254, 0); +lean_dec_ref(x_254); +x_259 = lean_ctor_get(x_255, 0); lean_inc(x_259); -x_260 = lean_ctor_get(x_254, 1); +x_260 = lean_ctor_get(x_255, 1); lean_inc_ref(x_260); -lean_dec_ref(x_254); +lean_dec_ref(x_255); x_261 = lean_int_mul(x_257, x_259); x_262 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__15; x_263 = lean_box(0); @@ -15316,14 +15316,14 @@ x_271 = l_Int_toNat(x_270); lean_dec(x_270); x_272 = l_Lean_instToExprInt_mkNat(x_271); x_273 = l_Lean_mkApp3(x_267, x_268, x_269, x_272); -x_87 = x_263; -x_88 = x_259; -x_89 = lean_box(0); -x_90 = x_262; -x_91 = x_264; -x_92 = x_260; -x_93 = x_258; -x_94 = x_261; +x_87 = x_258; +x_88 = x_264; +x_89 = x_261; +x_90 = x_263; +x_91 = x_259; +x_92 = lean_box(0); +x_93 = x_260; +x_94 = x_262; x_95 = x_273; goto block_113; } @@ -15333,14 +15333,14 @@ lean_object* x_274; lean_object* x_275; x_274 = l_Int_toNat(x_257); lean_dec(x_257); x_275 = l_Lean_instToExprInt_mkNat(x_274); -x_87 = x_263; -x_88 = x_259; -x_89 = lean_box(0); -x_90 = x_262; -x_91 = x_264; -x_92 = x_260; -x_93 = x_258; -x_94 = x_261; +x_87 = x_258; +x_88 = x_264; +x_89 = x_261; +x_90 = x_263; +x_91 = x_259; +x_92 = lean_box(0); +x_93 = x_260; +x_94 = x_262; x_95 = x_275; goto block_113; } @@ -15349,18 +15349,18 @@ case 1: { 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; uint8_t x_286; lean_dec(x_44); -x_276 = lean_ctor_get(x_255, 0); +x_276 = lean_ctor_get(x_254, 0); lean_inc(x_276); -x_277 = lean_ctor_get(x_255, 1); +x_277 = lean_ctor_get(x_254, 1); lean_inc_ref(x_277); -lean_dec_ref(x_255); -x_278 = lean_ctor_get(x_254, 0); +lean_dec_ref(x_254); +x_278 = lean_ctor_get(x_255, 0); lean_inc(x_278); -x_279 = lean_ctor_get(x_254, 1); +x_279 = lean_ctor_get(x_255, 1); lean_inc_ref(x_279); -x_280 = lean_ctor_get(x_254, 2); +x_280 = lean_ctor_get(x_255, 2); lean_inc_ref(x_280); -lean_dec_ref(x_254); +lean_dec_ref(x_255); x_281 = lean_int_mul(x_276, x_278); x_282 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__15; x_283 = lean_box(0); @@ -15379,14 +15379,14 @@ x_291 = l_Int_toNat(x_290); lean_dec(x_290); x_292 = l_Lean_instToExprInt_mkNat(x_291); x_293 = l_Lean_mkApp3(x_287, x_288, x_289, x_292); -x_156 = x_281; -x_157 = x_284; -x_158 = lean_box(0); -x_159 = x_282; -x_160 = x_277; -x_161 = x_278; -x_162 = x_280; -x_163 = x_279; +x_156 = x_282; +x_157 = x_277; +x_158 = x_279; +x_159 = x_284; +x_160 = x_280; +x_161 = lean_box(0); +x_162 = x_281; +x_163 = x_278; x_164 = x_283; x_165 = x_293; goto block_183; @@ -15397,14 +15397,14 @@ lean_object* x_294; lean_object* x_295; x_294 = l_Int_toNat(x_276); lean_dec(x_276); x_295 = l_Lean_instToExprInt_mkNat(x_294); -x_156 = x_281; -x_157 = x_284; -x_158 = lean_box(0); -x_159 = x_282; -x_160 = x_277; -x_161 = x_278; -x_162 = x_280; -x_163 = x_279; +x_156 = x_282; +x_157 = x_277; +x_158 = x_279; +x_159 = x_284; +x_160 = x_280; +x_161 = lean_box(0); +x_162 = x_281; +x_163 = x_278; x_164 = x_283; x_165 = x_295; goto block_183; @@ -15412,7 +15412,7 @@ goto block_183; } default: { -lean_dec_ref(x_255); +lean_dec_ref(x_254); lean_dec_ref(x_46); lean_dec_ref(x_45); lean_dec(x_44); @@ -15424,21 +15424,21 @@ goto block_18; case 1: { lean_dec(x_44); -if (lean_obj_tag(x_254) == 0) +if (lean_obj_tag(x_255) == 0) { lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -x_296 = lean_ctor_get(x_255, 0); +x_296 = lean_ctor_get(x_254, 0); lean_inc(x_296); -x_297 = lean_ctor_get(x_255, 1); +x_297 = lean_ctor_get(x_254, 1); lean_inc_ref(x_297); -x_298 = lean_ctor_get(x_255, 2); +x_298 = lean_ctor_get(x_254, 2); lean_inc_ref(x_298); -lean_dec_ref(x_255); -x_299 = lean_ctor_get(x_254, 0); +lean_dec_ref(x_254); +x_299 = lean_ctor_get(x_255, 0); lean_inc(x_299); -x_300 = lean_ctor_get(x_254, 1); +x_300 = lean_ctor_get(x_255, 1); lean_inc_ref(x_300); -lean_dec_ref(x_254); +lean_dec_ref(x_255); x_301 = lean_int_mul(x_296, x_299); x_302 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__15; x_303 = lean_box(0); @@ -15457,15 +15457,15 @@ x_311 = l_Int_toNat(x_310); lean_dec(x_310); x_312 = l_Lean_instToExprInt_mkNat(x_311); x_313 = l_Lean_mkApp3(x_307, x_308, x_309, x_312); -x_226 = lean_box(0); -x_227 = x_300; -x_228 = x_304; -x_229 = x_298; -x_230 = x_297; -x_231 = x_301; -x_232 = x_302; -x_233 = x_299; -x_234 = x_303; +x_226 = x_304; +x_227 = x_301; +x_228 = x_298; +x_229 = x_302; +x_230 = x_303; +x_231 = x_299; +x_232 = lean_box(0); +x_233 = x_300; +x_234 = x_297; x_235 = x_313; goto block_253; } @@ -15475,23 +15475,23 @@ lean_object* x_314; lean_object* x_315; x_314 = l_Int_toNat(x_296); lean_dec(x_296); x_315 = l_Lean_instToExprInt_mkNat(x_314); -x_226 = lean_box(0); -x_227 = x_300; -x_228 = x_304; -x_229 = x_298; -x_230 = x_297; -x_231 = x_301; -x_232 = x_302; -x_233 = x_299; -x_234 = x_303; +x_226 = x_304; +x_227 = x_301; +x_228 = x_298; +x_229 = x_302; +x_230 = x_303; +x_231 = x_299; +x_232 = lean_box(0); +x_233 = x_300; +x_234 = x_297; x_235 = x_315; goto block_253; } } else { -lean_dec_ref(x_255); -lean_dec(x_254); +lean_dec(x_255); +lean_dec_ref(x_254); lean_dec_ref(x_46); lean_dec_ref(x_45); x_15 = lean_box(0); @@ -15500,7 +15500,7 @@ goto block_18; } default: { -lean_dec(x_254); +lean_dec(x_255); lean_dec_ref(x_46); lean_dec_ref(x_45); lean_dec(x_44); @@ -15532,8 +15532,8 @@ x_335 = lean_int_dec_eq(x_332, x_334); if (x_335 == 0) { lean_free_object(x_329); -x_254 = x_331; -x_255 = x_317; +x_254 = x_317; +x_255 = x_331; x_256 = lean_box(0); goto block_316; } @@ -15574,8 +15574,8 @@ return x_329; else { lean_free_object(x_329); -x_254 = x_331; -x_255 = x_317; +x_254 = x_317; +x_255 = x_331; x_256 = lean_box(0); goto block_316; } @@ -15595,8 +15595,8 @@ x_347 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; x_348 = lean_int_dec_eq(x_345, x_347); if (x_348 == 0) { -x_254 = x_344; -x_255 = x_317; +x_254 = x_317; +x_255 = x_344; x_256 = lean_box(0); goto block_316; } @@ -15630,8 +15630,8 @@ return x_353; } else { -x_254 = x_344; -x_255 = x_317; +x_254 = x_317; +x_255 = x_344; x_256 = lean_box(0); goto block_316; } @@ -16627,12 +16627,12 @@ x_60 = l_Int_toNat(x_59); lean_dec(x_59); x_61 = l_Lean_instToExprInt_mkNat(x_60); x_62 = l_Lean_mkApp3(x_56, x_57, x_58, x_61); -x_16 = lean_box(0); -x_17 = x_53; -x_18 = x_45; -x_19 = x_52; +x_16 = x_53; +x_17 = x_47; +x_18 = x_52; +x_19 = x_45; x_20 = x_49; -x_21 = x_47; +x_21 = lean_box(0); x_22 = x_62; goto block_26; } @@ -16642,12 +16642,12 @@ lean_object* x_63; lean_object* x_64; x_63 = l_Int_toNat(x_44); lean_dec(x_44); x_64 = l_Lean_instToExprInt_mkNat(x_63); -x_16 = lean_box(0); -x_17 = x_53; -x_18 = x_45; -x_19 = x_52; +x_16 = x_53; +x_17 = x_47; +x_18 = x_52; +x_19 = x_45; x_20 = x_49; -x_21 = x_47; +x_21 = lean_box(0); x_22 = x_64; goto block_26; } @@ -16733,13 +16733,13 @@ x_90 = l_Int_toNat(x_89); lean_dec(x_89); x_91 = l_Lean_instToExprInt_mkNat(x_90); x_92 = l_Lean_mkApp3(x_86, x_87, x_88, x_91); -x_27 = x_67; -x_28 = x_77; -x_29 = x_75; -x_30 = lean_box(0); -x_31 = x_79; -x_32 = x_82; -x_33 = x_83; +x_27 = x_82; +x_28 = lean_box(0); +x_29 = x_67; +x_30 = x_75; +x_31 = x_83; +x_32 = x_79; +x_33 = x_77; x_34 = x_92; goto block_38; } @@ -16749,13 +16749,13 @@ lean_object* x_93; lean_object* x_94; x_93 = l_Int_toNat(x_65); lean_dec(x_65); x_94 = l_Lean_instToExprInt_mkNat(x_93); -x_27 = x_67; -x_28 = x_77; -x_29 = x_75; -x_30 = lean_box(0); -x_31 = x_79; -x_32 = x_82; -x_33 = x_83; +x_27 = x_82; +x_28 = lean_box(0); +x_29 = x_67; +x_30 = x_75; +x_31 = x_83; +x_32 = x_79; +x_33 = x_77; x_34 = x_94; goto block_38; } @@ -16884,7 +16884,7 @@ return x_111; { lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_24 = l_Lean_mkApp6(x_17, x_21, x_20, x_22, x_19, x_23, x_18); +x_24 = l_Lean_mkApp6(x_16, x_17, x_20, x_22, x_18, x_23, x_19); x_25 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_25, 0, x_24); return x_25; @@ -16893,7 +16893,7 @@ return x_25; { lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_36 = l_Lean_mkApp7(x_33, x_29, x_28, x_34, x_31, x_32, x_35, x_27); +x_36 = l_Lean_mkApp7(x_31, x_30, x_33, x_34, x_32, x_27, x_35, x_29); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); return x_37; @@ -17298,17 +17298,17 @@ x_199 = l_Int_toNat(x_198); lean_dec(x_198); x_200 = l_Lean_instToExprInt_mkNat(x_199); x_201 = l_Lean_mkApp3(x_195, x_196, x_197, x_200); -x_79 = x_190; +x_79 = x_187; x_80 = x_182; -x_81 = x_192; -x_82 = x_185; -x_83 = x_180; -x_84 = x_187; -x_85 = x_176; -x_86 = lean_box(0); -x_87 = x_191; -x_88 = x_189; -x_89 = x_188; +x_81 = x_176; +x_82 = x_192; +x_83 = x_189; +x_84 = x_191; +x_85 = x_188; +x_86 = x_190; +x_87 = x_185; +x_88 = x_180; +x_89 = lean_box(0); x_90 = x_201; goto block_150; } @@ -17317,17 +17317,17 @@ else lean_object* x_202; lean_object* x_203; x_202 = l_Int_toNat(x_1); x_203 = l_Lean_instToExprInt_mkNat(x_202); -x_79 = x_190; +x_79 = x_187; x_80 = x_182; -x_81 = x_192; -x_82 = x_185; -x_83 = x_180; -x_84 = x_187; -x_85 = x_176; -x_86 = lean_box(0); -x_87 = x_191; -x_88 = x_189; -x_89 = x_188; +x_81 = x_176; +x_82 = x_192; +x_83 = x_189; +x_84 = x_191; +x_85 = x_188; +x_86 = x_190; +x_87 = x_185; +x_88 = x_180; +x_89 = lean_box(0); x_90 = x_203; goto block_150; } @@ -17457,13 +17457,13 @@ 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; x_36 = lean_ctor_get(x_34, 0); -lean_inc_ref(x_20); +lean_inc_ref(x_27); lean_inc_ref(x_29); -x_37 = l_Lean_mkApp6(x_22, x_28, x_24, x_21, x_29, x_23, x_20); +x_37 = l_Lean_mkApp6(x_21, x_25, x_20, x_23, x_29, x_19, x_27); x_38 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_39 = l_Lean_Name_mkStr3(x_27, x_19, x_38); -x_40 = l_Lean_mkConst(x_39, x_26); -x_41 = l_Lean_mkApp6(x_40, x_31, x_33, x_29, x_36, x_20, x_37); +x_39 = l_Lean_Name_mkStr3(x_22, x_26, x_38); +x_40 = l_Lean_mkConst(x_39, x_24); +x_41 = l_Lean_mkApp6(x_40, x_31, x_33, x_29, x_36, x_27, x_37); lean_ctor_set(x_34, 0, x_41); return x_34; } @@ -17473,13 +17473,13 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean x_42 = lean_ctor_get(x_34, 0); lean_inc(x_42); lean_dec(x_34); -lean_inc_ref(x_20); +lean_inc_ref(x_27); lean_inc_ref(x_29); -x_43 = l_Lean_mkApp6(x_22, x_28, x_24, x_21, x_29, x_23, x_20); +x_43 = l_Lean_mkApp6(x_21, x_25, x_20, x_23, x_29, x_19, x_27); x_44 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_45 = l_Lean_Name_mkStr3(x_27, x_19, x_44); -x_46 = l_Lean_mkConst(x_45, x_26); -x_47 = l_Lean_mkApp6(x_46, x_31, x_33, x_29, x_42, x_20, x_43); +x_45 = l_Lean_Name_mkStr3(x_22, x_26, x_44); +x_46 = l_Lean_mkConst(x_45, x_24); +x_47 = l_Lean_mkApp6(x_46, x_31, x_33, x_29, x_42, x_27, x_43); x_48 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_48, 0, x_47); return x_48; @@ -17489,61 +17489,61 @@ return x_48; { lean_object* x_61; uint8_t x_62; x_61 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_62 = lean_int_dec_le(x_61, x_52); +x_62 = lean_int_dec_le(x_61, x_53); if (x_62 == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; x_63 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_64 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_58); +lean_inc(x_55); x_65 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_58); +lean_ctor_set(x_65, 1, x_55); x_66 = l_Lean_Expr_const___override(x_63, x_65); -lean_inc_ref(x_57); -x_67 = l_Lean_Name_mkStr1(x_57); -lean_inc(x_58); -x_68 = l_Lean_Expr_const___override(x_67, x_58); +lean_inc_ref(x_54); +x_67 = l_Lean_Name_mkStr1(x_54); +lean_inc(x_55); +x_68 = l_Lean_Expr_const___override(x_67, x_55); x_69 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_57); -x_70 = l_Lean_Name_mkStr2(x_57, x_69); -lean_inc(x_58); -x_71 = l_Lean_Expr_const___override(x_70, x_58); -x_72 = lean_int_neg(x_52); -lean_dec(x_52); +lean_inc_ref(x_54); +x_70 = l_Lean_Name_mkStr2(x_54, x_69); +lean_inc(x_55); +x_71 = l_Lean_Expr_const___override(x_70, x_55); +x_72 = lean_int_neg(x_53); +lean_dec(x_53); x_73 = l_Int_toNat(x_72); lean_dec(x_72); x_74 = l_Lean_instToExprInt_mkNat(x_73); x_75 = l_Lean_mkApp3(x_66, x_68, x_71, x_74); x_19 = x_51; x_20 = x_50; -x_21 = x_60; -x_22 = x_53; -x_23 = x_54; -x_24 = x_56; -x_25 = lean_box(0); -x_26 = x_58; -x_27 = x_57; -x_28 = x_59; +x_21 = x_52; +x_22 = x_54; +x_23 = x_60; +x_24 = x_55; +x_25 = x_56; +x_26 = x_57; +x_27 = x_59; +x_28 = lean_box(0); x_29 = x_75; goto block_49; } else { lean_object* x_76; lean_object* x_77; -x_76 = l_Int_toNat(x_52); -lean_dec(x_52); +x_76 = l_Int_toNat(x_53); +lean_dec(x_53); x_77 = l_Lean_instToExprInt_mkNat(x_76); x_19 = x_51; x_20 = x_50; -x_21 = x_60; -x_22 = x_53; -x_23 = x_54; -x_24 = x_56; -x_25 = lean_box(0); -x_26 = x_58; -x_27 = x_57; -x_28 = x_59; +x_21 = x_52; +x_22 = x_54; +x_23 = x_60; +x_24 = x_55; +x_25 = x_56; +x_26 = x_57; +x_27 = x_59; +x_28 = lean_box(0); x_29 = x_77; goto block_49; } @@ -17553,7 +17553,7 @@ goto block_49; lean_object* x_91; lean_object* x_92; x_91 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; lean_inc_ref(x_90); -x_92 = l_Lean_mkApp6(x_81, x_83, x_80, x_90, x_82, x_91, x_84); +x_92 = l_Lean_mkApp6(x_82, x_88, x_80, x_90, x_87, x_91, x_79); if (lean_obj_tag(x_2) == 0) { lean_object* x_93; @@ -17562,8 +17562,8 @@ lean_inc(x_14); lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); -lean_inc_ref(x_89); -x_93 = l_Lean_Meta_getIntValue_x3f(x_89, x_11, x_12, x_13, x_14); +lean_inc_ref(x_85); +x_93 = l_Lean_Meta_getIntValue_x3f(x_85, x_11, x_12, x_13, x_14); if (lean_obj_tag(x_93) == 0) { lean_object* x_94; @@ -17574,11 +17574,11 @@ if (lean_obj_tag(x_94) == 0) { lean_object* x_95; lean_object* x_96; lean_dec_ref(x_92); -lean_dec_ref(x_89); -lean_dec_ref(x_88); -lean_dec(x_87); +lean_dec_ref(x_86); lean_dec_ref(x_85); -lean_dec_ref(x_79); +lean_dec(x_84); +lean_dec_ref(x_83); +lean_dec_ref(x_81); lean_dec_ref(x_16); x_95 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__2; x_96 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_spec__0(x_95, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); @@ -17593,11 +17593,11 @@ lean_dec_ref(x_94); x_98 = lean_int_ediv(x_97, x_1); lean_dec(x_97); x_99 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__3; -lean_inc_ref(x_79); -lean_inc_ref(x_88); -x_100 = l_Lean_Name_mkStr3(x_88, x_79, x_99); -lean_inc(x_87); -x_101 = l_Lean_mkConst(x_100, x_87); +lean_inc_ref(x_86); +lean_inc_ref(x_83); +x_100 = l_Lean_Name_mkStr3(x_83, x_86, x_99); +lean_inc(x_84); +x_101 = l_Lean_mkConst(x_100, x_84); x_102 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; x_103 = lean_int_dec_le(x_102, x_1); if (x_103 == 0) @@ -17605,35 +17605,35 @@ 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_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; x_104 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_105 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_87); +lean_inc(x_84); x_106 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_87); +lean_ctor_set(x_106, 1, x_84); x_107 = l_Lean_Expr_const___override(x_104, x_106); -lean_inc_ref(x_88); -x_108 = l_Lean_Name_mkStr1(x_88); -lean_inc(x_87); -x_109 = l_Lean_Expr_const___override(x_108, x_87); +lean_inc_ref(x_83); +x_108 = l_Lean_Name_mkStr1(x_83); +lean_inc(x_84); +x_109 = l_Lean_Expr_const___override(x_108, x_84); x_110 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_88); -x_111 = l_Lean_Name_mkStr2(x_88, x_110); -lean_inc(x_87); -x_112 = l_Lean_Expr_const___override(x_111, x_87); +lean_inc_ref(x_83); +x_111 = l_Lean_Name_mkStr2(x_83, x_110); +lean_inc(x_84); +x_112 = l_Lean_Expr_const___override(x_111, x_84); x_113 = lean_int_neg(x_1); x_114 = l_Int_toNat(x_113); lean_dec(x_113); x_115 = l_Lean_instToExprInt_mkNat(x_114); x_116 = l_Lean_mkApp3(x_107, x_109, x_112, x_115); -x_50 = x_91; -x_51 = x_79; -x_52 = x_98; -x_53 = x_101; -x_54 = x_92; -x_55 = lean_box(0); +x_50 = x_81; +x_51 = x_92; +x_52 = x_101; +x_53 = x_98; +x_54 = x_83; +x_55 = x_84; x_56 = x_85; -x_57 = x_88; -x_58 = x_87; -x_59 = x_89; +x_57 = x_86; +x_58 = lean_box(0); +x_59 = x_91; x_60 = x_116; goto block_78; } @@ -17642,16 +17642,16 @@ else lean_object* x_117; lean_object* x_118; x_117 = l_Int_toNat(x_1); x_118 = l_Lean_instToExprInt_mkNat(x_117); -x_50 = x_91; -x_51 = x_79; -x_52 = x_98; -x_53 = x_101; -x_54 = x_92; -x_55 = lean_box(0); +x_50 = x_81; +x_51 = x_92; +x_52 = x_101; +x_53 = x_98; +x_54 = x_83; +x_55 = x_84; x_56 = x_85; -x_57 = x_88; -x_58 = x_87; -x_59 = x_89; +x_57 = x_86; +x_58 = lean_box(0); +x_59 = x_91; x_60 = x_118; goto block_78; } @@ -17661,11 +17661,11 @@ else { uint8_t x_119; lean_dec_ref(x_92); -lean_dec_ref(x_89); -lean_dec_ref(x_88); -lean_dec(x_87); +lean_dec_ref(x_86); lean_dec_ref(x_85); -lean_dec_ref(x_79); +lean_dec(x_84); +lean_dec_ref(x_83); +lean_dec_ref(x_81); lean_dec_ref(x_16); lean_dec(x_14); lean_dec_ref(x_13); @@ -17730,15 +17730,15 @@ if (x_130 == 0) lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; x_131 = lean_ctor_get(x_129, 0); x_132 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__4; -lean_inc_ref(x_79); -lean_inc_ref(x_88); -x_133 = l_Lean_Name_mkStr3(x_88, x_79, x_132); -lean_inc(x_87); -x_134 = l_Lean_mkConst(x_133, x_87); -x_135 = l_Lean_mkApp4(x_134, x_89, x_85, x_90, x_92); +lean_inc_ref(x_86); +lean_inc_ref(x_83); +x_133 = l_Lean_Name_mkStr3(x_83, x_86, x_132); +lean_inc(x_84); +x_134 = l_Lean_mkConst(x_133, x_84); +x_135 = l_Lean_mkApp4(x_134, x_85, x_81, x_90, x_92); x_136 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__5; -x_137 = l_Lean_Name_mkStr3(x_88, x_79, x_136); -x_138 = l_Lean_mkConst(x_137, x_87); +x_137 = l_Lean_Name_mkStr3(x_83, x_86, x_136); +x_138 = l_Lean_mkConst(x_137, x_84); x_139 = l_Lean_mkApp6(x_138, x_124, x_126, x_128, x_131, x_91, x_135); lean_ctor_set(x_129, 0, x_139); return x_129; @@ -17750,15 +17750,15 @@ x_140 = lean_ctor_get(x_129, 0); lean_inc(x_140); lean_dec(x_129); x_141 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__4; -lean_inc_ref(x_79); -lean_inc_ref(x_88); -x_142 = l_Lean_Name_mkStr3(x_88, x_79, x_141); -lean_inc(x_87); -x_143 = l_Lean_mkConst(x_142, x_87); -x_144 = l_Lean_mkApp4(x_143, x_89, x_85, x_90, x_92); +lean_inc_ref(x_86); +lean_inc_ref(x_83); +x_142 = l_Lean_Name_mkStr3(x_83, x_86, x_141); +lean_inc(x_84); +x_143 = l_Lean_mkConst(x_142, x_84); +x_144 = l_Lean_mkApp4(x_143, x_85, x_81, x_90, x_92); x_145 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__5; -x_146 = l_Lean_Name_mkStr3(x_88, x_79, x_145); -x_147 = l_Lean_mkConst(x_146, x_87); +x_146 = l_Lean_Name_mkStr3(x_83, x_86, x_145); +x_147 = l_Lean_mkConst(x_146, x_84); x_148 = l_Lean_mkApp6(x_147, x_124, x_126, x_128, x_140, x_91, x_144); x_149 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_149, 0, x_148); @@ -18142,17 +18142,17 @@ x_199 = l_Int_toNat(x_198); lean_dec(x_198); x_200 = l_Lean_instToExprInt_mkNat(x_199); x_201 = l_Lean_mkApp3(x_195, x_196, x_197, x_200); -x_79 = x_176; -x_80 = x_189; -x_81 = x_187; -x_82 = x_182; -x_83 = x_191; -x_84 = x_185; -x_85 = x_180; -x_86 = x_190; -x_87 = x_188; -x_88 = x_192; -x_89 = lean_box(0); +x_79 = lean_box(0); +x_80 = x_187; +x_81 = x_185; +x_82 = x_190; +x_83 = x_180; +x_84 = x_192; +x_85 = x_188; +x_86 = x_182; +x_87 = x_176; +x_88 = x_191; +x_89 = x_189; x_90 = x_201; goto block_150; } @@ -18161,17 +18161,17 @@ else lean_object* x_202; lean_object* x_203; x_202 = l_Int_toNat(x_1); x_203 = l_Lean_instToExprInt_mkNat(x_202); -x_79 = x_176; -x_80 = x_189; -x_81 = x_187; -x_82 = x_182; -x_83 = x_191; -x_84 = x_185; -x_85 = x_180; -x_86 = x_190; -x_87 = x_188; -x_88 = x_192; -x_89 = lean_box(0); +x_79 = lean_box(0); +x_80 = x_187; +x_81 = x_185; +x_82 = x_190; +x_83 = x_180; +x_84 = x_192; +x_85 = x_188; +x_86 = x_182; +x_87 = x_176; +x_88 = x_191; +x_89 = x_189; x_90 = x_203; goto block_150; } @@ -18301,13 +18301,13 @@ 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; x_36 = lean_ctor_get(x_34, 0); -lean_inc_ref(x_20); +lean_inc_ref(x_19); lean_inc_ref(x_29); -x_37 = l_Lean_mkApp6(x_24, x_28, x_21, x_26, x_29, x_22, x_20); +x_37 = l_Lean_mkApp6(x_28, x_22, x_25, x_24, x_29, x_23, x_19); x_38 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_39 = l_Lean_Name_mkStr3(x_19, x_27, x_38); -x_40 = l_Lean_mkConst(x_39, x_25); -x_41 = l_Lean_mkApp6(x_40, x_31, x_33, x_29, x_36, x_20, x_37); +x_39 = l_Lean_Name_mkStr3(x_27, x_20, x_38); +x_40 = l_Lean_mkConst(x_39, x_26); +x_41 = l_Lean_mkApp6(x_40, x_31, x_33, x_29, x_36, x_19, x_37); lean_ctor_set(x_34, 0, x_41); return x_34; } @@ -18317,13 +18317,13 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean x_42 = lean_ctor_get(x_34, 0); lean_inc(x_42); lean_dec(x_34); -lean_inc_ref(x_20); +lean_inc_ref(x_19); lean_inc_ref(x_29); -x_43 = l_Lean_mkApp6(x_24, x_28, x_21, x_26, x_29, x_22, x_20); +x_43 = l_Lean_mkApp6(x_28, x_22, x_25, x_24, x_29, x_23, x_19); x_44 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_45 = l_Lean_Name_mkStr3(x_19, x_27, x_44); -x_46 = l_Lean_mkConst(x_45, x_25); -x_47 = l_Lean_mkApp6(x_46, x_31, x_33, x_29, x_42, x_20, x_43); +x_45 = l_Lean_Name_mkStr3(x_27, x_20, x_44); +x_46 = l_Lean_mkConst(x_45, x_26); +x_47 = l_Lean_mkApp6(x_46, x_31, x_33, x_29, x_42, x_19, x_43); x_48 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_48, 0, x_47); return x_48; @@ -18333,7 +18333,7 @@ return x_48; { lean_object* x_61; uint8_t x_62; x_61 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_62 = lean_int_dec_le(x_61, x_55); +x_62 = lean_int_dec_le(x_61, x_58); if (x_62 == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; @@ -18344,50 +18344,50 @@ x_65 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_65, 0, x_64); lean_ctor_set(x_65, 1, x_57); x_66 = l_Lean_Expr_const___override(x_63, x_65); -lean_inc_ref(x_52); -x_67 = l_Lean_Name_mkStr1(x_52); +lean_inc_ref(x_56); +x_67 = l_Lean_Name_mkStr1(x_56); lean_inc(x_57); x_68 = l_Lean_Expr_const___override(x_67, x_57); x_69 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_52); -x_70 = l_Lean_Name_mkStr2(x_52, x_69); +lean_inc_ref(x_56); +x_70 = l_Lean_Name_mkStr2(x_56, x_69); lean_inc(x_57); x_71 = l_Lean_Expr_const___override(x_70, x_57); -x_72 = lean_int_neg(x_55); -lean_dec(x_55); +x_72 = lean_int_neg(x_58); +lean_dec(x_58); x_73 = l_Int_toNat(x_72); lean_dec(x_72); x_74 = l_Lean_instToExprInt_mkNat(x_73); x_75 = l_Lean_mkApp3(x_66, x_68, x_71, x_74); -x_19 = x_52; -x_20 = x_51; -x_21 = x_50; -x_22 = x_54; -x_23 = lean_box(0); -x_24 = x_56; -x_25 = x_57; -x_26 = x_60; -x_27 = x_59; -x_28 = x_58; +x_19 = x_50; +x_20 = x_52; +x_21 = lean_box(0); +x_22 = x_53; +x_23 = x_54; +x_24 = x_60; +x_25 = x_55; +x_26 = x_57; +x_27 = x_56; +x_28 = x_59; x_29 = x_75; goto block_49; } else { lean_object* x_76; lean_object* x_77; -x_76 = l_Int_toNat(x_55); -lean_dec(x_55); +x_76 = l_Int_toNat(x_58); +lean_dec(x_58); x_77 = l_Lean_instToExprInt_mkNat(x_76); -x_19 = x_52; -x_20 = x_51; -x_21 = x_50; -x_22 = x_54; -x_23 = lean_box(0); -x_24 = x_56; -x_25 = x_57; -x_26 = x_60; -x_27 = x_59; -x_28 = x_58; +x_19 = x_50; +x_20 = x_52; +x_21 = lean_box(0); +x_22 = x_53; +x_23 = x_54; +x_24 = x_60; +x_25 = x_55; +x_26 = x_57; +x_27 = x_56; +x_28 = x_59; x_29 = x_77; goto block_49; } @@ -18397,7 +18397,7 @@ goto block_49; lean_object* x_91; lean_object* x_92; x_91 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; lean_inc_ref(x_90); -x_92 = l_Lean_mkApp6(x_88, x_85, x_82, x_90, x_84, x_91, x_81); +x_92 = l_Lean_mkApp6(x_84, x_83, x_86, x_90, x_81, x_91, x_80); if (lean_obj_tag(x_2) == 0) { lean_object* x_93; @@ -18406,8 +18406,8 @@ lean_inc(x_14); lean_inc_ref(x_13); lean_inc(x_12); lean_inc_ref(x_11); -lean_inc_ref(x_87); -x_93 = l_Lean_Meta_getIntValue_x3f(x_87, x_11, x_12, x_13, x_14); +lean_inc_ref(x_85); +x_93 = l_Lean_Meta_getIntValue_x3f(x_85, x_11, x_12, x_13, x_14); if (lean_obj_tag(x_93) == 0) { lean_object* x_94; @@ -18418,11 +18418,11 @@ if (lean_obj_tag(x_94) == 0) { lean_object* x_95; lean_object* x_96; lean_dec_ref(x_92); +lean_dec_ref(x_89); +lean_dec(x_88); lean_dec_ref(x_87); -lean_dec_ref(x_86); -lean_dec(x_83); -lean_dec_ref(x_80); -lean_dec_ref(x_79); +lean_dec_ref(x_85); +lean_dec_ref(x_82); lean_dec_ref(x_16); x_95 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkModEqProof___closed__1; x_96 = l_panic___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_spec__0(x_95, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); @@ -18437,11 +18437,11 @@ lean_dec_ref(x_94); x_98 = lean_int_emod(x_97, x_1); lean_dec(x_97); x_99 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkModEqProof___closed__2; -lean_inc_ref(x_86); -lean_inc_ref(x_80); -x_100 = l_Lean_Name_mkStr3(x_80, x_86, x_99); -lean_inc(x_83); -x_101 = l_Lean_mkConst(x_100, x_83); +lean_inc_ref(x_82); +lean_inc_ref(x_89); +x_100 = l_Lean_Name_mkStr3(x_89, x_82, x_99); +lean_inc(x_88); +x_101 = l_Lean_mkConst(x_100, x_88); x_102 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; x_103 = lean_int_dec_le(x_102, x_1); if (x_103 == 0) @@ -18449,35 +18449,35 @@ 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_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; x_104 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_105 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_83); +lean_inc(x_88); x_106 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_83); +lean_ctor_set(x_106, 1, x_88); x_107 = l_Lean_Expr_const___override(x_104, x_106); -lean_inc_ref(x_80); -x_108 = l_Lean_Name_mkStr1(x_80); -lean_inc(x_83); -x_109 = l_Lean_Expr_const___override(x_108, x_83); +lean_inc_ref(x_89); +x_108 = l_Lean_Name_mkStr1(x_89); +lean_inc(x_88); +x_109 = l_Lean_Expr_const___override(x_108, x_88); x_110 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_80); -x_111 = l_Lean_Name_mkStr2(x_80, x_110); -lean_inc(x_83); -x_112 = l_Lean_Expr_const___override(x_111, x_83); +lean_inc_ref(x_89); +x_111 = l_Lean_Name_mkStr2(x_89, x_110); +lean_inc(x_88); +x_112 = l_Lean_Expr_const___override(x_111, x_88); x_113 = lean_int_neg(x_1); x_114 = l_Int_toNat(x_113); lean_dec(x_113); x_115 = l_Lean_instToExprInt_mkNat(x_114); x_116 = l_Lean_mkApp3(x_107, x_109, x_112, x_115); -x_50 = x_79; -x_51 = x_91; -x_52 = x_80; -x_53 = lean_box(0); +x_50 = x_91; +x_51 = lean_box(0); +x_52 = x_82; +x_53 = x_85; x_54 = x_92; -x_55 = x_98; -x_56 = x_101; -x_57 = x_83; -x_58 = x_87; -x_59 = x_86; +x_55 = x_87; +x_56 = x_89; +x_57 = x_88; +x_58 = x_98; +x_59 = x_101; x_60 = x_116; goto block_78; } @@ -18486,16 +18486,16 @@ else lean_object* x_117; lean_object* x_118; x_117 = l_Int_toNat(x_1); x_118 = l_Lean_instToExprInt_mkNat(x_117); -x_50 = x_79; -x_51 = x_91; -x_52 = x_80; -x_53 = lean_box(0); +x_50 = x_91; +x_51 = lean_box(0); +x_52 = x_82; +x_53 = x_85; x_54 = x_92; -x_55 = x_98; -x_56 = x_101; -x_57 = x_83; -x_58 = x_87; -x_59 = x_86; +x_55 = x_87; +x_56 = x_89; +x_57 = x_88; +x_58 = x_98; +x_59 = x_101; x_60 = x_118; goto block_78; } @@ -18505,11 +18505,11 @@ else { uint8_t x_119; lean_dec_ref(x_92); +lean_dec_ref(x_89); +lean_dec(x_88); lean_dec_ref(x_87); -lean_dec_ref(x_86); -lean_dec(x_83); -lean_dec_ref(x_80); -lean_dec_ref(x_79); +lean_dec_ref(x_85); +lean_dec_ref(x_82); lean_dec_ref(x_16); lean_dec(x_14); lean_dec_ref(x_13); @@ -18574,15 +18574,15 @@ if (x_130 == 0) lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; x_131 = lean_ctor_get(x_129, 0); x_132 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkModEqProof___closed__3; -lean_inc_ref(x_86); -lean_inc_ref(x_80); -x_133 = l_Lean_Name_mkStr3(x_80, x_86, x_132); -lean_inc(x_83); -x_134 = l_Lean_mkConst(x_133, x_83); -x_135 = l_Lean_mkApp4(x_134, x_87, x_79, x_90, x_92); +lean_inc_ref(x_82); +lean_inc_ref(x_89); +x_133 = l_Lean_Name_mkStr3(x_89, x_82, x_132); +lean_inc(x_88); +x_134 = l_Lean_mkConst(x_133, x_88); +x_135 = l_Lean_mkApp4(x_134, x_85, x_87, x_90, x_92); x_136 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__5; -x_137 = l_Lean_Name_mkStr3(x_80, x_86, x_136); -x_138 = l_Lean_mkConst(x_137, x_83); +x_137 = l_Lean_Name_mkStr3(x_89, x_82, x_136); +x_138 = l_Lean_mkConst(x_137, x_88); x_139 = l_Lean_mkApp6(x_138, x_124, x_126, x_128, x_131, x_91, x_135); lean_ctor_set(x_129, 0, x_139); return x_129; @@ -18594,15 +18594,15 @@ x_140 = lean_ctor_get(x_129, 0); lean_inc(x_140); lean_dec(x_129); x_141 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkModEqProof___closed__3; -lean_inc_ref(x_86); -lean_inc_ref(x_80); -x_142 = l_Lean_Name_mkStr3(x_80, x_86, x_141); -lean_inc(x_83); -x_143 = l_Lean_mkConst(x_142, x_83); -x_144 = l_Lean_mkApp4(x_143, x_87, x_79, x_90, x_92); +lean_inc_ref(x_82); +lean_inc_ref(x_89); +x_142 = l_Lean_Name_mkStr3(x_89, x_82, x_141); +lean_inc(x_88); +x_143 = l_Lean_mkConst(x_142, x_88); +x_144 = l_Lean_mkApp4(x_143, x_85, x_87, x_90, x_92); x_145 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkDivEqProof___closed__5; -x_146 = l_Lean_Name_mkStr3(x_80, x_86, x_145); -x_147 = l_Lean_mkConst(x_146, x_83); +x_146 = l_Lean_Name_mkStr3(x_89, x_82, x_145); +x_147 = l_Lean_mkConst(x_146, x_88); x_148 = l_Lean_mkApp6(x_147, x_124, x_126, x_128, x_140, x_91, x_144); x_149 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_149, 0, x_148); @@ -19083,13 +19083,13 @@ lean_dec(x_317); x_319 = l_Lean_instToExprInt_mkNat(x_318); x_320 = l_Lean_mkApp3(x_314, x_315, x_316, x_319); x_254 = x_310; -x_255 = x_305; +x_255 = x_308; x_256 = x_311; x_257 = x_298; x_258 = x_297; -x_259 = x_308; +x_259 = lean_box(0); x_260 = x_301; -x_261 = lean_box(0); +x_261 = x_305; x_262 = x_320; goto block_265; } @@ -19099,13 +19099,13 @@ lean_object* x_321; lean_object* x_322; x_321 = l_Int_toNat(x_1); x_322 = l_Lean_instToExprInt_mkNat(x_321); x_254 = x_310; -x_255 = x_305; +x_255 = x_308; x_256 = x_311; x_257 = x_298; x_258 = x_297; -x_259 = x_308; +x_259 = lean_box(0); x_260 = x_301; -x_261 = lean_box(0); +x_261 = x_305; x_262 = x_322; goto block_265; } @@ -19215,26 +19215,26 @@ return x_334; block_63: { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_42 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_28); +x_42 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_33); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); lean_dec_ref(x_42); lean_inc(x_19); -x_44 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkVarDecl(x_19, x_28, x_38, x_32, x_39, x_37, x_26, x_30, x_22, x_24, x_31); +x_44 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkVarDecl(x_19, x_33, x_28, x_24, x_21, x_40, x_30, x_23, x_27, x_25, x_31); x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); lean_dec_ref(x_44); -x_46 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_17, x_28, x_38, x_32, x_39, x_37, x_26, x_30, x_22, x_24, x_31); +x_46 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_17, x_33, x_28, x_24, x_21, x_40, x_30, x_23, x_27, x_25, x_31); lean_dec(x_31); -lean_dec_ref(x_24); -lean_dec(x_22); -lean_dec_ref(x_30); -lean_dec(x_26); -lean_dec_ref(x_37); -lean_dec(x_39); -lean_dec(x_32); -lean_dec(x_38); -lean_dec_ref(x_28); +lean_dec_ref(x_25); +lean_dec(x_27); +lean_dec_ref(x_23); +lean_dec(x_30); +lean_dec_ref(x_40); +lean_dec(x_21); +lean_dec(x_24); +lean_dec(x_28); +lean_dec_ref(x_33); x_47 = !lean_is_exclusive(x_46); if (x_47 == 0) { @@ -19242,10 +19242,10 @@ lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean x_48 = lean_ctor_get(x_46, 0); x_49 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; lean_inc_ref(x_41); -x_50 = l_Lean_mkApp8(x_33, x_34, x_36, x_23, x_40, x_41, x_20, x_27, x_49); +x_50 = l_Lean_mkApp8(x_26, x_22, x_35, x_32, x_36, x_41, x_37, x_20, x_49); x_51 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_52 = l_Lean_Name_mkStr3(x_21, x_25, x_51); -x_53 = l_Lean_mkConst(x_52, x_35); +x_52 = l_Lean_Name_mkStr3(x_39, x_29, x_51); +x_53 = l_Lean_mkConst(x_52, x_34); x_54 = l_Lean_mkApp6(x_53, x_43, x_45, x_41, x_48, x_49, x_50); lean_ctor_set(x_46, 0, x_54); return x_46; @@ -19258,10 +19258,10 @@ lean_inc(x_55); lean_dec(x_46); x_56 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; lean_inc_ref(x_41); -x_57 = l_Lean_mkApp8(x_33, x_34, x_36, x_23, x_40, x_41, x_20, x_27, x_56); +x_57 = l_Lean_mkApp8(x_26, x_22, x_35, x_32, x_36, x_41, x_37, x_20, x_56); x_58 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof___closed__0; -x_59 = l_Lean_Name_mkStr3(x_21, x_25, x_58); -x_60 = l_Lean_mkConst(x_59, x_35); +x_59 = l_Lean_Name_mkStr3(x_39, x_29, x_58); +x_60 = l_Lean_mkConst(x_59, x_34); x_61 = l_Lean_mkApp6(x_60, x_43, x_45, x_41, x_55, x_56, x_57); x_62 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_62, 0, x_61); @@ -19272,83 +19272,83 @@ return x_62; { lean_object* x_86; uint8_t x_87; x_86 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_87 = lean_int_dec_le(x_86, x_65); +x_87 = lean_int_dec_le(x_86, x_79); if (x_87 == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; x_88 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_89 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_80); +lean_inc(x_78); x_90 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_80); +lean_ctor_set(x_90, 1, x_78); x_91 = l_Lean_Expr_const___override(x_88, x_90); -lean_inc_ref(x_66); -x_92 = l_Lean_Name_mkStr1(x_66); -lean_inc(x_80); -x_93 = l_Lean_Expr_const___override(x_92, x_80); +lean_inc_ref(x_83); +x_92 = l_Lean_Name_mkStr1(x_83); +lean_inc(x_78); +x_93 = l_Lean_Expr_const___override(x_92, x_78); x_94 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_66); -x_95 = l_Lean_Name_mkStr2(x_66, x_94); -lean_inc(x_80); -x_96 = l_Lean_Expr_const___override(x_95, x_80); -x_97 = lean_int_neg(x_65); -lean_dec(x_65); +lean_inc_ref(x_83); +x_95 = l_Lean_Name_mkStr2(x_83, x_94); +lean_inc(x_78); +x_96 = l_Lean_Expr_const___override(x_95, x_78); +x_97 = lean_int_neg(x_79); +lean_dec(x_79); x_98 = l_Int_toNat(x_97); lean_dec(x_97); x_99 = l_Lean_instToExprInt_mkNat(x_98); x_100 = l_Lean_mkApp3(x_91, x_93, x_96, x_99); x_20 = x_64; -x_21 = x_66; -x_22 = x_67; -x_23 = x_68; -x_24 = x_69; -x_25 = x_70; -x_26 = x_71; -x_27 = x_72; -x_28 = x_73; -x_29 = lean_box(0); -x_30 = x_75; -x_31 = x_76; -x_32 = x_78; +x_21 = x_65; +x_22 = x_66; +x_23 = x_67; +x_24 = x_68; +x_25 = x_69; +x_26 = x_70; +x_27 = x_71; +x_28 = x_72; +x_29 = x_73; +x_30 = x_74; +x_31 = x_75; +x_32 = x_76; x_33 = x_77; -x_34 = x_79; +x_34 = x_78; x_35 = x_80; -x_36 = x_82; -x_37 = x_81; -x_38 = x_84; +x_36 = x_85; +x_37 = x_82; +x_38 = lean_box(0); x_39 = x_83; -x_40 = x_85; +x_40 = x_84; x_41 = x_100; goto block_63; } else { lean_object* x_101; lean_object* x_102; -x_101 = l_Int_toNat(x_65); -lean_dec(x_65); +x_101 = l_Int_toNat(x_79); +lean_dec(x_79); x_102 = l_Lean_instToExprInt_mkNat(x_101); x_20 = x_64; -x_21 = x_66; -x_22 = x_67; -x_23 = x_68; -x_24 = x_69; -x_25 = x_70; -x_26 = x_71; -x_27 = x_72; -x_28 = x_73; -x_29 = lean_box(0); -x_30 = x_75; -x_31 = x_76; -x_32 = x_78; +x_21 = x_65; +x_22 = x_66; +x_23 = x_67; +x_24 = x_68; +x_25 = x_69; +x_26 = x_70; +x_27 = x_71; +x_28 = x_72; +x_29 = x_73; +x_30 = x_74; +x_31 = x_75; +x_32 = x_76; x_33 = x_77; -x_34 = x_79; +x_34 = x_78; x_35 = x_80; -x_36 = x_82; -x_37 = x_81; -x_38 = x_84; +x_36 = x_85; +x_37 = x_82; +x_38 = lean_box(0); x_39 = x_83; -x_40 = x_85; +x_40 = x_84; x_41 = x_102; goto block_63; } @@ -19357,83 +19357,83 @@ goto block_63; { lean_object* x_126; uint8_t x_127; x_126 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_127 = lean_int_dec_le(x_126, x_107); +x_127 = lean_int_dec_le(x_126, x_104); if (x_127 == 0) { lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; x_128 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_129 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_120); +lean_inc(x_118); x_130 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_120); +lean_ctor_set(x_130, 1, x_118); x_131 = l_Lean_Expr_const___override(x_128, x_130); -lean_inc_ref(x_106); -x_132 = l_Lean_Name_mkStr1(x_106); -lean_inc(x_120); -x_133 = l_Lean_Expr_const___override(x_132, x_120); +lean_inc_ref(x_123); +x_132 = l_Lean_Name_mkStr1(x_123); +lean_inc(x_118); +x_133 = l_Lean_Expr_const___override(x_132, x_118); x_134 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_106); -x_135 = l_Lean_Name_mkStr2(x_106, x_134); -lean_inc(x_120); -x_136 = l_Lean_Expr_const___override(x_135, x_120); -x_137 = lean_int_neg(x_107); -lean_dec(x_107); +lean_inc_ref(x_123); +x_135 = l_Lean_Name_mkStr2(x_123, x_134); +lean_inc(x_118); +x_136 = l_Lean_Expr_const___override(x_135, x_118); +x_137 = lean_int_neg(x_104); +lean_dec(x_104); x_138 = l_Int_toNat(x_137); lean_dec(x_137); x_139 = l_Lean_instToExprInt_mkNat(x_138); x_140 = l_Lean_mkApp3(x_131, x_133, x_136, x_139); -x_64 = x_104; -x_65 = x_105; -x_66 = x_106; +x_64 = x_105; +x_65 = x_106; +x_66 = x_107; x_67 = x_108; -x_68 = x_125; -x_69 = x_109; -x_70 = x_110; -x_71 = x_111; -x_72 = x_112; -x_73 = x_113; -x_74 = lean_box(0); -x_75 = x_115; -x_76 = x_116; -x_77 = x_118; -x_78 = x_117; +x_68 = x_109; +x_69 = x_110; +x_70 = x_111; +x_71 = x_112; +x_72 = x_113; +x_73 = x_114; +x_74 = x_115; +x_75 = x_116; +x_76 = x_125; +x_77 = x_117; +x_78 = x_118; x_79 = x_119; x_80 = x_120; -x_81 = x_122; +x_81 = lean_box(0); x_82 = x_121; -x_83 = x_124; -x_84 = x_123; +x_83 = x_123; +x_84 = x_124; x_85 = x_140; goto block_103; } else { lean_object* x_141; lean_object* x_142; -x_141 = l_Int_toNat(x_107); -lean_dec(x_107); +x_141 = l_Int_toNat(x_104); +lean_dec(x_104); x_142 = l_Lean_instToExprInt_mkNat(x_141); -x_64 = x_104; -x_65 = x_105; -x_66 = x_106; +x_64 = x_105; +x_65 = x_106; +x_66 = x_107; x_67 = x_108; -x_68 = x_125; -x_69 = x_109; -x_70 = x_110; -x_71 = x_111; -x_72 = x_112; -x_73 = x_113; -x_74 = lean_box(0); -x_75 = x_115; -x_76 = x_116; -x_77 = x_118; -x_78 = x_117; +x_68 = x_109; +x_69 = x_110; +x_70 = x_111; +x_71 = x_112; +x_72 = x_113; +x_73 = x_114; +x_74 = x_115; +x_75 = x_116; +x_76 = x_125; +x_77 = x_117; +x_78 = x_118; x_79 = x_119; x_80 = x_120; -x_81 = x_122; +x_81 = lean_box(0); x_82 = x_121; -x_83 = x_124; -x_84 = x_123; +x_83 = x_123; +x_84 = x_124; x_85 = x_142; goto block_103; } @@ -19461,26 +19461,26 @@ lean_dec(x_170); x_172 = l_Lean_instToExprInt_mkNat(x_171); x_173 = l_Lean_mkApp3(x_167, x_168, x_169, x_172); x_104 = x_144; -x_105 = x_160; -x_106 = x_161; +x_105 = x_148; +x_106 = x_152; x_107 = x_145; -x_108 = x_156; -x_109 = x_157; -x_110 = x_162; -x_111 = x_154; -x_112 = x_148; -x_113 = x_149; -x_114 = lean_box(0); -x_115 = x_155; +x_108 = x_155; +x_109 = x_151; +x_110 = x_157; +x_111 = x_164; +x_112 = x_156; +x_113 = x_150; +x_114 = x_162; +x_115 = x_154; x_116 = x_158; -x_117 = x_151; -x_118 = x_164; -x_119 = x_146; -x_120 = x_163; +x_117 = x_149; +x_118 = x_163; +x_119 = x_160; +x_120 = x_146; x_121 = x_147; -x_122 = x_153; -x_123 = x_150; -x_124 = x_152; +x_122 = lean_box(0); +x_123 = x_161; +x_124 = x_153; x_125 = x_173; goto block_143; } @@ -19490,26 +19490,26 @@ lean_object* x_174; lean_object* x_175; x_174 = l_Int_toNat(x_1); x_175 = l_Lean_instToExprInt_mkNat(x_174); x_104 = x_144; -x_105 = x_160; -x_106 = x_161; +x_105 = x_148; +x_106 = x_152; x_107 = x_145; -x_108 = x_156; -x_109 = x_157; -x_110 = x_162; -x_111 = x_154; -x_112 = x_148; -x_113 = x_149; -x_114 = lean_box(0); -x_115 = x_155; +x_108 = x_155; +x_109 = x_151; +x_110 = x_157; +x_111 = x_164; +x_112 = x_156; +x_113 = x_150; +x_114 = x_162; +x_115 = x_154; x_116 = x_158; -x_117 = x_151; -x_118 = x_164; -x_119 = x_146; -x_120 = x_163; +x_117 = x_149; +x_118 = x_163; +x_119 = x_160; +x_120 = x_146; x_121 = x_147; -x_122 = x_153; -x_123 = x_150; -x_124 = x_152; +x_122 = lean_box(0); +x_123 = x_161; +x_124 = x_153; x_125 = x_175; goto block_143; } @@ -19518,22 +19518,22 @@ goto block_143; { lean_object* x_198; lean_object* x_199; x_198 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_199 = l_Lean_mkApp6(x_185, x_182, x_181, x_197, x_186, x_198, x_195); +x_199 = l_Lean_mkApp6(x_190, x_183, x_179, x_197, x_192, x_198, x_194); x_144 = x_177; -x_145 = x_179; -x_146 = x_190; -x_147 = x_194; +x_145 = x_181; +x_146 = x_193; +x_147 = x_196; x_148 = x_199; -x_149 = x_191; -x_150 = x_192; -x_151 = x_187; -x_152 = x_180; -x_153 = x_193; -x_154 = x_188; -x_155 = x_183; -x_156 = x_196; -x_157 = x_189; -x_158 = x_184; +x_149 = x_187; +x_150 = x_195; +x_151 = x_188; +x_152 = x_191; +x_153 = x_186; +x_154 = x_182; +x_155 = x_189; +x_156 = x_184; +x_157 = x_185; +x_158 = x_180; x_159 = lean_box(0); goto block_176; } @@ -19549,10 +19549,10 @@ x_216 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind x_217 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkContext___redArg___closed__8; x_218 = l_Lean_mkIntLit(x_215); x_219 = l_Lean_mkAppB(x_216, x_217, x_218); -x_144 = x_203; -x_145 = x_215; -x_146 = x_201; -x_147 = x_202; +x_144 = x_215; +x_145 = x_201; +x_146 = x_202; +x_147 = x_203; x_148 = x_219; x_149 = x_204; x_150 = x_205; @@ -19644,26 +19644,26 @@ x_242 = l_Int_toNat(x_241); lean_dec(x_241); x_243 = l_Lean_instToExprInt_mkNat(x_242); x_244 = l_Lean_mkApp3(x_238, x_239, x_240, x_243); -x_177 = x_203; +x_177 = x_215; x_178 = lean_box(0); -x_179 = x_215; -x_180 = x_207; -x_181 = x_229; -x_182 = x_225; -x_183 = x_210; -x_184 = x_213; -x_185 = x_235; -x_186 = x_232; -x_187 = x_206; -x_188 = x_209; -x_189 = x_212; -x_190 = x_201; -x_191 = x_204; -x_192 = x_205; -x_193 = x_208; -x_194 = x_202; -x_195 = x_234; -x_196 = x_211; +x_179 = x_229; +x_180 = x_213; +x_181 = x_201; +x_182 = x_209; +x_183 = x_225; +x_184 = x_211; +x_185 = x_212; +x_186 = x_208; +x_187 = x_204; +x_188 = x_206; +x_189 = x_210; +x_190 = x_235; +x_191 = x_207; +x_192 = x_232; +x_193 = x_202; +x_194 = x_234; +x_195 = x_205; +x_196 = x_203; x_197 = x_244; goto block_200; } @@ -19672,26 +19672,26 @@ else lean_object* x_245; lean_object* x_246; x_245 = l_Int_toNat(x_215); x_246 = l_Lean_instToExprInt_mkNat(x_245); -x_177 = x_203; +x_177 = x_215; x_178 = lean_box(0); -x_179 = x_215; -x_180 = x_207; -x_181 = x_229; -x_182 = x_225; -x_183 = x_210; -x_184 = x_213; -x_185 = x_235; -x_186 = x_232; -x_187 = x_206; -x_188 = x_209; -x_189 = x_212; -x_190 = x_201; -x_191 = x_204; -x_192 = x_205; -x_193 = x_208; -x_194 = x_202; -x_195 = x_234; -x_196 = x_211; +x_179 = x_229; +x_180 = x_213; +x_181 = x_201; +x_182 = x_209; +x_183 = x_225; +x_184 = x_211; +x_185 = x_212; +x_186 = x_208; +x_187 = x_204; +x_188 = x_206; +x_189 = x_210; +x_190 = x_235; +x_191 = x_207; +x_192 = x_232; +x_193 = x_202; +x_194 = x_234; +x_195 = x_205; +x_196 = x_203; x_197 = x_246; goto block_200; } @@ -19800,7 +19800,7 @@ return x_252; { lean_object* x_263; lean_object* x_264; x_263 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_264 = l_Lean_mkApp6(x_256, x_260, x_255, x_262, x_259, x_263, x_254); +x_264 = l_Lean_mkApp6(x_256, x_260, x_261, x_262, x_255, x_263, x_254); x_201 = x_257; x_202 = x_258; x_203 = x_264; @@ -21531,7 +21531,7 @@ return x_3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___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) { _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; lean_object* x_28; lean_object* x_29; 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_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_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_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_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_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_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_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_262; lean_object* x_263; lean_object* x_264; uint8_t 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_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_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_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_686; lean_object* x_687; uint8_t x_688; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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_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_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_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_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_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_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_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; 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; 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_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_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_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_686; lean_object* x_687; uint8_t x_688; lean_inc(x_1); x_686 = l_Lean_isTracingEnabledFor___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof_spec__0___redArg(x_1, x_11); x_687 = lean_ctor_get(x_686, 0); @@ -21624,11 +21624,11 @@ x_32 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_ x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); lean_dec_ref(x_32); -x_34 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkExprDecl(x_16, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28); +x_34 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkExprDecl(x_17, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); lean_dec_ref(x_34); -x_36 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_17, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28); +x_36 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_14, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28); lean_dec(x_28); lean_dec_ref(x_27); lean_dec(x_26); @@ -21644,7 +21644,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = lean_ctor_get(x_36, 0); -x_39 = l_Lean_Expr_app___override(x_14, x_18); +x_39 = l_Lean_Expr_app___override(x_16, x_18); x_40 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__2; x_41 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; x_42 = l_Lean_mkApp6(x_40, x_31, x_33, x_35, x_38, x_41, x_39); @@ -21657,7 +21657,7 @@ lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean x_43 = lean_ctor_get(x_36, 0); lean_inc(x_43); lean_dec(x_36); -x_44 = l_Lean_Expr_app___override(x_14, x_18); +x_44 = l_Lean_Expr_app___override(x_16, x_18); x_45 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__2; x_46 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; x_47 = l_Lean_mkApp6(x_45, x_31, x_33, x_35, x_43, x_46, x_44); @@ -21670,7 +21670,7 @@ return x_48; { lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_58 = l_Lean_mkApp6(x_53, x_52, x_54, x_55, x_56, x_57, x_51); +x_58 = l_Lean_mkApp6(x_51, x_53, x_52, x_54, x_56, x_57, x_55); x_59 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_59, 0, x_58); return x_59; @@ -21679,7 +21679,7 @@ return x_59; { lean_object* x_70; lean_object* x_71; lean_object* x_72; x_70 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_71 = l_Lean_mkApp8(x_68, x_66, x_65, x_62, x_63, x_69, x_70, x_67, x_64); +x_71 = l_Lean_mkApp8(x_68, x_66, x_62, x_67, x_65, x_69, x_70, x_64, x_63); x_72 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_72, 0, x_71); return x_72; @@ -21687,46 +21687,46 @@ return x_72; block_113: { lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_91 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_78); +x_91 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_79); x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); lean_dec_ref(x_91); -x_93 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkVarDecl(x_83, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +x_93 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkVarDecl(x_84, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); x_94 = lean_ctor_get(x_93, 0); lean_inc(x_94); lean_dec_ref(x_93); -x_95 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_75, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +x_95 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_77, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); x_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); lean_dec_ref(x_95); -x_97 = lean_ctor_get(x_84, 0); +x_97 = lean_ctor_get(x_75, 0); lean_inc_ref(x_97); -x_98 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_97, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +x_98 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_97, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); lean_dec_ref(x_98); -x_100 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_79, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +x_100 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_78, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); lean_dec_ref(x_100); -lean_inc(x_85); -lean_inc_ref(x_82); -lean_inc(x_80); -lean_inc_ref(x_81); +lean_inc(x_86); +lean_inc_ref(x_80); +lean_inc(x_89); +lean_inc_ref(x_85); +lean_inc(x_76); +lean_inc_ref(x_88); lean_inc(x_74); -lean_inc_ref(x_89); -lean_inc(x_77); -lean_inc(x_88); -lean_inc(x_87); -lean_inc_ref(x_78); -x_102 = lean_cutsat_eq_cnstr_to_proof(x_86, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +lean_inc(x_81); +lean_inc(x_82); +lean_inc_ref(x_79); +x_102 = lean_cutsat_eq_cnstr_to_proof(x_83, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); if (lean_obj_tag(x_102) == 0) { lean_object* x_103; lean_object* x_104; x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); lean_dec_ref(x_102); -x_104 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_84, x_78, x_87, x_88, x_77, x_89, x_74, x_81, x_80, x_82, x_85); +x_104 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_75, x_79, x_82, x_81, x_74, x_88, x_76, x_85, x_89, x_80, x_86); if (lean_obj_tag(x_104) == 0) { uint8_t x_105; @@ -21773,16 +21773,16 @@ lean_dec(x_96); lean_dec(x_94); lean_dec(x_92); lean_dec_ref(x_90); -lean_dec_ref(x_89); -lean_dec(x_88); -lean_dec(x_87); -lean_dec(x_85); -lean_dec_ref(x_84); -lean_dec_ref(x_82); -lean_dec_ref(x_81); -lean_dec(x_80); -lean_dec_ref(x_78); -lean_dec(x_77); +lean_dec(x_89); +lean_dec_ref(x_88); +lean_dec(x_86); +lean_dec_ref(x_85); +lean_dec(x_82); +lean_dec(x_81); +lean_dec_ref(x_80); +lean_dec_ref(x_79); +lean_dec(x_76); +lean_dec_ref(x_75); lean_dec(x_74); return x_102; } @@ -21791,7 +21791,7 @@ return x_102; { lean_object* x_123; lean_object* x_124; lean_object* x_125; x_123 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_124 = l_Lean_mkApp8(x_119, x_118, x_114, x_116, x_115, x_122, x_117, x_121, x_123); +x_124 = l_Lean_mkApp8(x_120, x_121, x_117, x_119, x_115, x_122, x_118, x_114, x_123); x_125 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_125, 0, x_124); return x_125; @@ -21799,23 +21799,23 @@ return x_125; block_169: { 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_145 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_132); +x_145 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_133); x_146 = lean_ctor_get(x_145, 0); lean_inc(x_146); lean_dec_ref(x_145); -x_147 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_127, x_132, x_140, x_142, x_130, x_143, x_128, x_135, x_134, x_136, x_139); +x_147 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_142, x_133, x_137, x_135, x_129, x_141, x_130, x_138, x_143, x_134, x_139); x_148 = lean_ctor_get(x_147, 0); lean_inc(x_148); lean_dec_ref(x_147); -x_149 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_141, x_132, x_140, x_142, x_130, x_143, x_128, x_135, x_134, x_136, x_139); +x_149 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_127, x_133, x_137, x_135, x_129, x_141, x_130, x_138, x_143, x_134, x_139); x_150 = lean_ctor_get(x_149, 0); lean_inc(x_150); lean_dec_ref(x_149); -x_151 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_133, x_132, x_140, x_142, x_130, x_143, x_128, x_135, x_134, x_136, x_139); +x_151 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_132, x_133, x_137, x_135, x_129, x_141, x_130, x_138, x_143, x_134, x_139); x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); lean_dec_ref(x_151); -x_153 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_131, x_132, x_140, x_142, x_130, x_143, x_128, x_135, x_134, x_136, x_139); +x_153 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_128, x_133, x_137, x_135, x_129, x_141, x_130, x_138, x_143, x_134, x_139); if (lean_obj_tag(x_153) == 0) { lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; @@ -21824,46 +21824,46 @@ lean_inc(x_154); lean_dec_ref(x_153); x_155 = lean_box(0); x_156 = l_Lean_mkConst(x_144, x_155); -x_157 = l_Lean_mkNatLit(x_138); +x_157 = l_Lean_mkNatLit(x_136); x_158 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_159 = lean_int_dec_le(x_158, x_137); +x_159 = lean_int_dec_le(x_158, x_131); if (x_159 == 0) { lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; x_160 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__10; x_161 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__12; x_162 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__15; -x_163 = lean_int_neg(x_137); -lean_dec(x_137); +x_163 = lean_int_neg(x_131); +lean_dec(x_131); x_164 = l_Int_toNat(x_163); lean_dec(x_163); x_165 = l_Lean_instToExprInt_mkNat(x_164); x_166 = l_Lean_mkApp3(x_160, x_161, x_162, x_165); -x_114 = x_148; +x_114 = x_154; x_115 = x_157; -x_116 = x_150; -x_117 = x_152; -x_118 = x_146; -x_119 = x_156; -x_120 = lean_box(0); -x_121 = x_154; +x_116 = lean_box(0); +x_117 = x_148; +x_118 = x_152; +x_119 = x_150; +x_120 = x_156; +x_121 = x_146; x_122 = x_166; goto block_126; } else { lean_object* x_167; lean_object* x_168; -x_167 = l_Int_toNat(x_137); -lean_dec(x_137); +x_167 = l_Int_toNat(x_131); +lean_dec(x_131); x_168 = l_Lean_instToExprInt_mkNat(x_167); -x_114 = x_148; +x_114 = x_154; x_115 = x_157; -x_116 = x_150; -x_117 = x_152; -x_118 = x_146; -x_119 = x_156; -x_120 = lean_box(0); -x_121 = x_154; +x_116 = lean_box(0); +x_117 = x_148; +x_118 = x_152; +x_119 = x_150; +x_120 = x_156; +x_121 = x_146; x_122 = x_168; goto block_126; } @@ -21875,8 +21875,8 @@ lean_dec(x_150); lean_dec(x_148); lean_dec(x_146); lean_dec(x_144); -lean_dec(x_138); -lean_dec(x_137); +lean_dec(x_136); +lean_dec(x_131); return x_153; } } @@ -21884,7 +21884,7 @@ return x_153; { lean_object* x_181; lean_object* x_182; lean_object* x_183; x_181 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_182 = l_Lean_mkApp10(x_179, x_174, x_170, x_173, x_178, x_177, x_171, x_180, x_175, x_172, x_181); +x_182 = l_Lean_mkApp10(x_176, x_171, x_179, x_175, x_170, x_177, x_174, x_180, x_173, x_172, x_181); x_183 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_183, 0, x_182); return x_183; @@ -21892,39 +21892,39 @@ return x_183; block_214: { lean_object* x_197; lean_object* x_198; uint8_t x_199; -x_197 = l_Lean_mkNatLit(x_189); +x_197 = l_Lean_mkNatLit(x_194); x_198 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_199 = lean_int_dec_le(x_198, x_187); +x_199 = lean_int_dec_le(x_198, x_191); if (x_199 == 0) { lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; x_200 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_201 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_186); +lean_inc(x_187); x_202 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_186); +lean_ctor_set(x_202, 1, x_187); x_203 = l_Lean_Expr_const___override(x_200, x_202); x_204 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__11; -lean_inc(x_186); -x_205 = l_Lean_Expr_const___override(x_204, x_186); +lean_inc(x_187); +x_205 = l_Lean_Expr_const___override(x_204, x_187); x_206 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__14; -x_207 = l_Lean_Expr_const___override(x_206, x_186); -x_208 = lean_int_neg(x_187); -lean_dec(x_187); +x_207 = l_Lean_Expr_const___override(x_206, x_187); +x_208 = lean_int_neg(x_191); +lean_dec(x_191); x_209 = l_Int_toNat(x_208); lean_dec(x_208); x_210 = l_Lean_instToExprInt_mkNat(x_209); x_211 = l_Lean_mkApp3(x_203, x_205, x_207, x_210); x_170 = x_185; -x_171 = x_197; -x_172 = x_188; -x_173 = x_191; -x_174 = x_190; -x_175 = x_192; -x_176 = lean_box(0); +x_171 = x_186; +x_172 = x_190; +x_173 = x_189; +x_174 = x_197; +x_175 = x_188; +x_176 = x_192; x_177 = x_196; -x_178 = x_193; +x_178 = lean_box(0); x_179 = x_195; x_180 = x_211; goto block_184; @@ -21932,19 +21932,19 @@ goto block_184; else { lean_object* x_212; lean_object* x_213; -lean_dec(x_186); -x_212 = l_Int_toNat(x_187); lean_dec(x_187); +x_212 = l_Int_toNat(x_191); +lean_dec(x_191); x_213 = l_Lean_instToExprInt_mkNat(x_212); x_170 = x_185; -x_171 = x_197; -x_172 = x_188; -x_173 = x_191; -x_174 = x_190; -x_175 = x_192; -x_176 = lean_box(0); +x_171 = x_186; +x_172 = x_190; +x_173 = x_189; +x_174 = x_197; +x_175 = x_188; +x_176 = x_192; x_177 = x_196; -x_178 = x_193; +x_178 = lean_box(0); x_179 = x_195; x_180 = x_213; goto block_184; @@ -21953,32 +21953,32 @@ goto block_184; block_261: { lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_234 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_220); +x_234 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_221); x_235 = lean_ctor_get(x_234, 0); lean_inc(x_235); lean_dec_ref(x_234); -x_236 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_215, x_220, x_229, x_231, x_218, x_232, x_216, x_224, x_222, x_225, x_228); +x_236 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_231, x_221, x_226, x_223, x_217, x_230, x_218, x_227, x_232, x_222, x_228); x_237 = lean_ctor_get(x_236, 0); lean_inc(x_237); lean_dec_ref(x_236); -x_238 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_230, x_220, x_229, x_231, x_218, x_232, x_216, x_224, x_222, x_225, x_228); +x_238 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_215, x_221, x_226, x_223, x_217, x_230, x_218, x_227, x_232, x_222, x_228); x_239 = lean_ctor_get(x_238, 0); lean_inc(x_239); lean_dec_ref(x_238); -x_240 = lean_ctor_get(x_223, 0); +x_240 = lean_ctor_get(x_225, 0); lean_inc(x_240); -x_241 = lean_ctor_get(x_223, 1); +x_241 = lean_ctor_get(x_225, 1); lean_inc_ref(x_241); -lean_dec_ref(x_223); -x_242 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_241, x_220, x_229, x_231, x_218, x_232, x_216, x_224, x_222, x_225, x_228); +lean_dec_ref(x_225); +x_242 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_241, x_221, x_226, x_223, x_217, x_230, x_218, x_227, x_232, x_222, x_228); x_243 = lean_ctor_get(x_242, 0); lean_inc(x_243); lean_dec_ref(x_242); -x_244 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_221, x_220, x_229, x_231, x_218, x_232, x_216, x_224, x_222, x_225, x_228); +x_244 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_220, x_221, x_226, x_223, x_217, x_230, x_218, x_227, x_232, x_222, x_228); x_245 = lean_ctor_get(x_244, 0); lean_inc(x_245); lean_dec_ref(x_244); -x_246 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_219, x_220, x_229, x_231, x_218, x_232, x_216, x_224, x_222, x_225, x_228); +x_246 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_216, x_221, x_226, x_223, x_217, x_230, x_218, x_227, x_232, x_222, x_228); if (lean_obj_tag(x_246) == 0) { lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; uint8_t x_251; @@ -22001,17 +22001,17 @@ x_256 = l_Int_toNat(x_255); lean_dec(x_255); x_257 = l_Lean_instToExprInt_mkNat(x_256); x_258 = l_Lean_mkApp3(x_252, x_253, x_254, x_257); -x_185 = x_237; -x_186 = x_248; -x_187 = x_226; -x_188 = x_247; -x_189 = x_227; -x_190 = x_235; -x_191 = x_239; -x_192 = x_245; -x_193 = x_243; -x_194 = lean_box(0); -x_195 = x_249; +x_185 = x_243; +x_186 = x_235; +x_187 = x_248; +x_188 = x_239; +x_189 = x_245; +x_190 = x_247; +x_191 = x_219; +x_192 = x_249; +x_193 = lean_box(0); +x_194 = x_224; +x_195 = x_237; x_196 = x_258; goto block_214; } @@ -22021,17 +22021,17 @@ lean_object* x_259; lean_object* x_260; x_259 = l_Int_toNat(x_240); lean_dec(x_240); x_260 = l_Lean_instToExprInt_mkNat(x_259); -x_185 = x_237; -x_186 = x_248; -x_187 = x_226; -x_188 = x_247; -x_189 = x_227; -x_190 = x_235; -x_191 = x_239; -x_192 = x_245; -x_193 = x_243; -x_194 = lean_box(0); -x_195 = x_249; +x_185 = x_243; +x_186 = x_235; +x_187 = x_248; +x_188 = x_239; +x_189 = x_245; +x_190 = x_247; +x_191 = x_219; +x_192 = x_249; +x_193 = lean_box(0); +x_194 = x_224; +x_195 = x_237; x_196 = x_260; goto block_214; } @@ -22045,34 +22045,34 @@ lean_dec(x_239); lean_dec(x_237); lean_dec(x_235); lean_dec(x_233); -lean_dec(x_227); -lean_dec(x_226); +lean_dec(x_224); +lean_dec(x_219); return x_246; } } block_287: { -if (lean_obj_tag(x_263) == 0) +if (lean_obj_tag(x_273) == 0) { -if (x_265 == 0) +if (x_272 == 0) { lean_object* x_281; x_281 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__4; x_127 = x_262; -x_128 = x_264; -x_129 = lean_box(0); -x_130 = x_267; -x_131 = x_268; -x_132 = x_269; -x_133 = x_270; -x_134 = x_271; -x_135 = x_272; -x_136 = x_273; -x_137 = x_280; +x_128 = x_263; +x_129 = x_264; +x_130 = x_265; +x_131 = x_280; +x_132 = x_266; +x_133 = x_267; +x_134 = x_268; +x_135 = x_269; +x_136 = x_270; +x_137 = x_271; x_138 = x_274; x_139 = x_275; -x_140 = x_277; -x_141 = x_276; +x_140 = lean_box(0); +x_141 = x_277; x_142 = x_279; x_143 = x_278; x_144 = x_281; @@ -22083,20 +22083,20 @@ else lean_object* x_282; x_282 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__6; x_127 = x_262; -x_128 = x_264; -x_129 = lean_box(0); -x_130 = x_267; -x_131 = x_268; -x_132 = x_269; -x_133 = x_270; -x_134 = x_271; -x_135 = x_272; -x_136 = x_273; -x_137 = x_280; +x_128 = x_263; +x_129 = x_264; +x_130 = x_265; +x_131 = x_280; +x_132 = x_266; +x_133 = x_267; +x_134 = x_268; +x_135 = x_269; +x_136 = x_270; +x_137 = x_271; x_138 = x_274; x_139 = x_275; -x_140 = x_277; -x_141 = x_276; +x_140 = lean_box(0); +x_141 = x_277; x_142 = x_279; x_143 = x_278; x_144 = x_282; @@ -22105,29 +22105,29 @@ goto block_169; } else { -if (x_265 == 0) +if (x_272 == 0) { lean_object* x_283; lean_object* x_284; -x_283 = lean_ctor_get(x_263, 0); +x_283 = lean_ctor_get(x_273, 0); lean_inc(x_283); -lean_dec_ref(x_263); +lean_dec_ref(x_273); x_284 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__8; x_215 = x_262; -x_216 = x_264; -x_217 = lean_box(0); -x_218 = x_267; -x_219 = x_268; -x_220 = x_269; -x_221 = x_270; -x_222 = x_271; -x_223 = x_283; -x_224 = x_272; -x_225 = x_273; -x_226 = x_280; +x_216 = x_263; +x_217 = x_264; +x_218 = x_265; +x_219 = x_280; +x_220 = x_266; +x_221 = x_267; +x_222 = x_268; +x_223 = x_269; +x_224 = x_270; +x_225 = x_283; +x_226 = x_271; x_227 = x_274; x_228 = x_275; -x_229 = x_277; -x_230 = x_276; +x_229 = lean_box(0); +x_230 = x_277; x_231 = x_279; x_232 = x_278; x_233 = x_284; @@ -22136,26 +22136,26 @@ goto block_261; else { lean_object* x_285; lean_object* x_286; -x_285 = lean_ctor_get(x_263, 0); +x_285 = lean_ctor_get(x_273, 0); lean_inc(x_285); -lean_dec_ref(x_263); +lean_dec_ref(x_273); x_286 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__10; x_215 = x_262; -x_216 = x_264; -x_217 = lean_box(0); -x_218 = x_267; -x_219 = x_268; -x_220 = x_269; -x_221 = x_270; -x_222 = x_271; -x_223 = x_285; -x_224 = x_272; -x_225 = x_273; -x_226 = x_280; +x_216 = x_263; +x_217 = x_264; +x_218 = x_265; +x_219 = x_280; +x_220 = x_266; +x_221 = x_267; +x_222 = x_268; +x_223 = x_269; +x_224 = x_270; +x_225 = x_285; +x_226 = x_271; x_227 = x_274; x_228 = x_275; -x_229 = x_277; -x_230 = x_276; +x_229 = lean_box(0); +x_230 = x_277; x_231 = x_279; x_232 = x_278; x_233 = x_286; @@ -22167,7 +22167,7 @@ goto block_261; { lean_object* x_297; lean_object* x_298; lean_object* x_299; x_297 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_298 = l_Lean_mkApp8(x_294, x_288, x_296, x_290, x_295, x_291, x_297, x_293, x_289); +x_298 = l_Lean_mkApp8(x_292, x_290, x_296, x_288, x_291, x_295, x_297, x_294, x_289); x_299 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_299, 0, x_298); return x_299; @@ -22176,7 +22176,7 @@ return x_299; { lean_object* x_310; lean_object* x_311; lean_object* x_312; x_310 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_311 = l_Lean_mkApp8(x_308, x_302, x_309, x_304, x_303, x_305, x_310, x_307, x_306); +x_311 = l_Lean_mkApp8(x_304, x_308, x_309, x_307, x_301, x_306, x_310, x_305, x_302); x_312 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_312, 0, x_311); return x_312; @@ -22370,10 +22370,10 @@ x_359 = lean_ctor_get(x_358, 0); lean_inc(x_359); lean_dec_ref(x_358); x_360 = l_Lean_Meta_mkOfEqFalseCore(x_354, x_359); -x_14 = x_355; +x_14 = x_353; x_15 = x_356; -x_16 = x_357; -x_17 = x_353; +x_16 = x_355; +x_17 = x_357; x_18 = x_360; x_19 = x_314; x_20 = x_315; @@ -22440,10 +22440,10 @@ x_367 = lean_ctor_get(x_366, 0); lean_inc(x_367); lean_dec_ref(x_366); x_368 = l_Lean_Meta_mkOfEqTrueCore(x_362, x_367); -x_14 = x_363; +x_14 = x_361; x_15 = x_364; -x_16 = x_365; -x_17 = x_361; +x_16 = x_363; +x_17 = x_365; x_18 = x_368; x_19 = x_314; x_20 = x_315; @@ -22694,11 +22694,11 @@ lean_dec(x_425); x_427 = l_Lean_instToExprInt_mkNat(x_426); x_428 = l_Lean_mkApp3(x_422, x_423, x_424, x_427); x_50 = lean_box(0); -x_51 = x_416; -x_52 = x_409; -x_53 = x_418; -x_54 = x_412; -x_55 = x_414; +x_51 = x_418; +x_52 = x_412; +x_53 = x_409; +x_54 = x_414; +x_55 = x_416; x_56 = x_428; goto block_60; } @@ -22709,11 +22709,11 @@ x_429 = l_Int_toNat(x_419); lean_dec(x_419); x_430 = l_Lean_instToExprInt_mkNat(x_429); x_50 = lean_box(0); -x_51 = x_416; -x_52 = x_409; -x_53 = x_418; -x_54 = x_412; -x_55 = x_414; +x_51 = x_418; +x_52 = x_412; +x_53 = x_409; +x_54 = x_414; +x_55 = x_416; x_56 = x_430; goto block_60; } @@ -22907,12 +22907,12 @@ lean_dec(x_481); x_483 = l_Lean_instToExprInt_mkNat(x_482); x_484 = l_Lean_mkApp3(x_478, x_479, x_480, x_483); x_61 = lean_box(0); -x_62 = x_468; -x_63 = x_470; -x_64 = x_474; -x_65 = x_465; +x_62 = x_465; +x_63 = x_474; +x_64 = x_472; +x_65 = x_470; x_66 = x_462; -x_67 = x_472; +x_67 = x_468; x_68 = x_475; x_69 = x_484; goto block_73; @@ -22924,12 +22924,12 @@ x_485 = l_Int_toNat(x_460); lean_dec(x_460); x_486 = l_Lean_instToExprInt_mkNat(x_485); x_61 = lean_box(0); -x_62 = x_468; -x_63 = x_470; -x_64 = x_474; -x_65 = x_465; +x_62 = x_465; +x_63 = x_474; +x_64 = x_472; +x_65 = x_470; x_66 = x_462; -x_67 = x_472; +x_67 = x_468; x_68 = x_475; x_69 = x_486; goto block_73; @@ -22990,22 +22990,22 @@ if (x_494 == 0) { lean_object* x_495; x_495 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__33; -x_74 = x_319; -x_75 = x_491; -x_76 = lean_box(0); -x_77 = x_317; -x_78 = x_314; -x_79 = x_488; -x_80 = x_321; -x_81 = x_320; -x_82 = x_322; -x_83 = x_489; -x_84 = x_490; -x_85 = x_323; -x_86 = x_487; -x_87 = x_315; -x_88 = x_316; -x_89 = x_318; +x_74 = x_317; +x_75 = x_490; +x_76 = x_319; +x_77 = x_491; +x_78 = x_488; +x_79 = x_314; +x_80 = x_322; +x_81 = x_316; +x_82 = x_315; +x_83 = x_487; +x_84 = x_489; +x_85 = x_320; +x_86 = x_323; +x_87 = lean_box(0); +x_88 = x_318; +x_89 = x_321; x_90 = x_495; goto block_113; } @@ -23013,22 +23013,22 @@ else { lean_object* x_496; x_496 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof___lam__0___closed__36; -x_74 = x_319; -x_75 = x_491; -x_76 = lean_box(0); -x_77 = x_317; -x_78 = x_314; -x_79 = x_488; -x_80 = x_321; -x_81 = x_320; -x_82 = x_322; -x_83 = x_489; -x_84 = x_490; -x_85 = x_323; -x_86 = x_487; -x_87 = x_315; -x_88 = x_316; -x_89 = x_318; +x_74 = x_317; +x_75 = x_490; +x_76 = x_319; +x_77 = x_491; +x_78 = x_488; +x_79 = x_314; +x_80 = x_322; +x_81 = x_316; +x_82 = x_315; +x_83 = x_487; +x_84 = x_489; +x_85 = x_320; +x_86 = x_323; +x_87 = lean_box(0); +x_88 = x_318; +x_89 = x_321; x_90 = x_496; goto block_113; } @@ -23325,24 +23325,24 @@ lean_inc_ref(x_579); x_580 = lean_ctor_get(x_575, 0); lean_inc_ref(x_580); x_581 = l_Int_Linear_Poly_leadCoeff(x_579); -x_262 = x_579; -x_263 = x_578; -x_264 = x_319; -x_265 = x_573; -x_266 = lean_box(0); -x_267 = x_317; -x_268 = x_571; -x_269 = x_314; -x_270 = x_576; -x_271 = x_321; -x_272 = x_320; -x_273 = x_322; -x_274 = x_577; +x_262 = x_580; +x_263 = x_571; +x_264 = x_317; +x_265 = x_319; +x_266 = x_576; +x_267 = x_314; +x_268 = x_322; +x_269 = x_316; +x_270 = x_577; +x_271 = x_315; +x_272 = x_573; +x_273 = x_578; +x_274 = x_320; x_275 = x_323; -x_276 = x_580; -x_277 = x_315; -x_278 = x_318; -x_279 = x_316; +x_276 = lean_box(0); +x_277 = x_318; +x_278 = x_321; +x_279 = x_579; x_280 = x_581; goto block_287; } @@ -23363,24 +23363,24 @@ lean_inc_ref(x_587); x_588 = lean_ctor_get(x_583, 0); lean_inc_ref(x_588); x_589 = l_Int_Linear_Poly_leadCoeff(x_588); -x_262 = x_587; -x_263 = x_586; -x_264 = x_319; -x_265 = x_573; -x_266 = lean_box(0); -x_267 = x_317; -x_268 = x_571; -x_269 = x_314; -x_270 = x_584; -x_271 = x_321; -x_272 = x_320; -x_273 = x_322; -x_274 = x_585; +x_262 = x_588; +x_263 = x_571; +x_264 = x_317; +x_265 = x_319; +x_266 = x_584; +x_267 = x_314; +x_268 = x_322; +x_269 = x_316; +x_270 = x_585; +x_271 = x_315; +x_272 = x_573; +x_273 = x_586; +x_274 = x_320; x_275 = x_323; -x_276 = x_588; -x_277 = x_315; -x_278 = x_318; -x_279 = x_316; +x_276 = lean_box(0); +x_277 = x_318; +x_278 = x_321; +x_279 = x_587; x_280 = x_589; goto block_287; } @@ -23457,14 +23457,14 @@ x_615 = l_Int_toNat(x_614); lean_dec(x_614); x_616 = l_Lean_instToExprInt_mkNat(x_615); x_617 = l_Lean_mkApp3(x_611, x_612, x_613, x_616); -x_288 = x_594; +x_288 = x_598; x_289 = x_607; -x_290 = x_598; -x_291 = x_603; -x_292 = lean_box(0); -x_293 = x_605; -x_294 = x_608; -x_295 = x_601; +x_290 = x_594; +x_291 = x_601; +x_292 = x_608; +x_293 = lean_box(0); +x_294 = x_605; +x_295 = x_603; x_296 = x_617; goto block_300; } @@ -23474,14 +23474,14 @@ lean_object* x_618; lean_object* x_619; x_618 = l_Int_toNat(x_595); lean_dec(x_595); x_619 = l_Lean_instToExprInt_mkNat(x_618); -x_288 = x_594; +x_288 = x_598; x_289 = x_607; -x_290 = x_598; -x_291 = x_603; -x_292 = lean_box(0); -x_293 = x_605; -x_294 = x_608; -x_295 = x_601; +x_290 = x_594; +x_291 = x_601; +x_292 = x_608; +x_293 = lean_box(0); +x_294 = x_605; +x_295 = x_603; x_296 = x_619; goto block_300; } @@ -23590,14 +23590,14 @@ x_645 = l_Int_toNat(x_644); lean_dec(x_644); x_646 = l_Lean_instToExprInt_mkNat(x_645); x_647 = l_Lean_mkApp3(x_641, x_642, x_643, x_646); -x_301 = lean_box(0); -x_302 = x_624; -x_303 = x_631; -x_304 = x_628; -x_305 = x_633; -x_306 = x_637; -x_307 = x_635; -x_308 = x_638; +x_301 = x_631; +x_302 = x_637; +x_303 = lean_box(0); +x_304 = x_638; +x_305 = x_635; +x_306 = x_633; +x_307 = x_628; +x_308 = x_624; x_309 = x_647; goto block_313; } @@ -23607,14 +23607,14 @@ lean_object* x_648; lean_object* x_649; x_648 = l_Int_toNat(x_625); lean_dec(x_625); x_649 = l_Lean_instToExprInt_mkNat(x_648); -x_301 = lean_box(0); -x_302 = x_624; -x_303 = x_631; -x_304 = x_628; -x_305 = x_633; -x_306 = x_637; -x_307 = x_635; -x_308 = x_638; +x_301 = x_631; +x_302 = x_637; +x_303 = lean_box(0); +x_304 = x_638; +x_305 = x_635; +x_306 = x_633; +x_307 = x_628; +x_308 = x_624; x_309 = x_649; goto block_313; } @@ -24448,7 +24448,7 @@ lean_ctor_set(x_250, 0, x_248); lean_ctor_set(x_250, 1, x_249); x_251 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_251, 0, x_250); -lean_ctor_set(x_251, 1, x_238); +lean_ctor_set(x_251, 1, x_237); x_252 = l_Lean_throwError___at___00__private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getVarOf_spec__3___redArg(x_251, x_9, x_10, x_11, x_12); lean_dec(x_12); lean_dec_ref(x_11); @@ -24501,8 +24501,8 @@ if (x_24 == 0) lean_object* x_272; x_272 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__31; x_236 = x_271; -x_237 = lean_box(0); -x_238 = x_257; +x_237 = x_257; +x_238 = lean_box(0); x_239 = x_272; goto block_256; } @@ -24511,8 +24511,8 @@ else lean_object* x_273; x_273 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__32; x_236 = x_271; -x_237 = lean_box(0); -x_238 = x_257; +x_237 = x_257; +x_238 = lean_box(0); x_239 = x_273; goto block_256; } @@ -24666,68 +24666,68 @@ return x_59; block_97: { 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_74 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_63); +x_74 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_65); x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); lean_dec_ref(x_74); lean_inc_ref(x_26); -x_76 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +x_76 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); x_77 = lean_ctor_get(x_76, 0); lean_inc(x_77); lean_dec_ref(x_76); lean_inc_ref(x_27); -x_78 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +x_78 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); lean_dec_ref(x_78); -lean_inc(x_70); -lean_inc_ref(x_71); -lean_inc(x_65); -lean_inc_ref(x_72); +lean_inc(x_61); +lean_inc_ref(x_70); lean_inc(x_69); lean_inc_ref(x_66); -lean_inc(x_61); +lean_inc(x_71); +lean_inc_ref(x_62); lean_inc(x_64); -lean_inc(x_67); -lean_inc_ref(x_63); -x_80 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_20, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +lean_inc(x_63); +lean_inc(x_72); +lean_inc_ref(x_65); +x_80 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_20, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); if (lean_obj_tag(x_80) == 0) { lean_object* x_81; lean_object* x_82; x_81 = lean_ctor_get(x_80, 0); lean_inc(x_81); lean_dec_ref(x_80); -lean_inc(x_70); -lean_inc_ref(x_71); -lean_inc(x_65); -lean_inc_ref(x_72); +lean_inc(x_61); +lean_inc_ref(x_70); lean_inc(x_69); lean_inc_ref(x_66); -lean_inc(x_61); +lean_inc(x_71); +lean_inc_ref(x_62); lean_inc(x_64); -lean_inc(x_67); -lean_inc_ref(x_63); -x_82 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_21, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +lean_inc(x_63); +lean_inc(x_72); +lean_inc_ref(x_65); +x_82 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_21, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); if (lean_obj_tag(x_82) == 0) { lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; x_83 = lean_ctor_get(x_82, 0); lean_inc(x_83); lean_dec_ref(x_82); -x_84 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_63); +x_84 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_65); x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); lean_dec_ref(x_84); -x_86 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +x_86 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); x_87 = lean_ctor_get(x_86, 0); lean_inc(x_87); lean_dec_ref(x_86); -x_88 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_63, x_67, x_64, x_61, x_66, x_69, x_72, x_65, x_71, x_70); +x_88 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_65, x_72, x_63, x_64, x_62, x_71, x_66, x_69, x_70, x_61); x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); lean_dec_ref(x_88); x_90 = lean_box(0); -x_91 = l_Lean_mkConst(x_62, x_90); +x_91 = l_Lean_mkConst(x_68, x_90); lean_inc(x_28); x_92 = l_Lean_mkNatLit(x_28); x_93 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; @@ -24736,16 +24736,16 @@ x_95 = l_Lean_mkConst(x_73, x_90); x_96 = l_Lean_mkApp3(x_95, x_85, x_87, x_89); x_29 = x_94; x_30 = x_96; -x_31 = x_63; -x_32 = x_67; -x_33 = x_64; -x_34 = x_61; -x_35 = x_66; -x_36 = x_69; -x_37 = x_72; -x_38 = x_65; -x_39 = x_71; -x_40 = x_70; +x_31 = x_65; +x_32 = x_72; +x_33 = x_63; +x_34 = x_64; +x_35 = x_62; +x_36 = x_71; +x_37 = x_66; +x_38 = x_69; +x_39 = x_70; +x_40 = x_61; x_41 = lean_box(0); goto block_60; } @@ -24756,16 +24756,16 @@ lean_dec(x_79); lean_dec(x_77); lean_dec(x_75); lean_dec(x_73); -lean_dec_ref(x_72); -lean_dec_ref(x_71); -lean_dec(x_70); +lean_dec(x_72); +lean_dec(x_71); +lean_dec_ref(x_70); lean_dec(x_69); -lean_dec(x_67); +lean_dec(x_68); lean_dec_ref(x_66); -lean_dec(x_65); +lean_dec_ref(x_65); lean_dec(x_64); -lean_dec_ref(x_63); -lean_dec(x_62); +lean_dec(x_63); +lean_dec_ref(x_62); lean_dec(x_61); lean_dec(x_28); lean_dec_ref(x_27); @@ -24781,16 +24781,16 @@ lean_dec(x_79); lean_dec(x_77); lean_dec(x_75); lean_dec(x_73); -lean_dec_ref(x_72); -lean_dec_ref(x_71); -lean_dec(x_70); +lean_dec(x_72); +lean_dec(x_71); +lean_dec_ref(x_70); lean_dec(x_69); -lean_dec(x_67); +lean_dec(x_68); lean_dec_ref(x_66); -lean_dec(x_65); +lean_dec_ref(x_65); lean_dec(x_64); -lean_dec_ref(x_63); -lean_dec(x_62); +lean_dec(x_63); +lean_dec_ref(x_62); lean_dec(x_61); lean_dec(x_28); lean_dec_ref(x_27); @@ -24808,17 +24808,17 @@ if (x_24 == 0) lean_object* x_110; x_110 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__4; x_61 = x_98; -x_62 = x_109; +x_62 = x_100; x_63 = x_99; -x_64 = x_100; -x_65 = x_102; -x_66 = x_101; -x_67 = x_103; -x_68 = lean_box(0); +x_64 = x_102; +x_65 = x_101; +x_66 = x_103; +x_67 = lean_box(0); +x_68 = x_109; x_69 = x_105; -x_70 = x_106; -x_71 = x_108; -x_72 = x_107; +x_70 = x_108; +x_71 = x_107; +x_72 = x_106; x_73 = x_110; goto block_97; } @@ -24827,17 +24827,17 @@ else lean_object* x_111; x_111 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__6; x_61 = x_98; -x_62 = x_109; +x_62 = x_100; x_63 = x_99; -x_64 = x_100; -x_65 = x_102; -x_66 = x_101; -x_67 = x_103; -x_68 = lean_box(0); +x_64 = x_102; +x_65 = x_101; +x_66 = x_103; +x_67 = lean_box(0); +x_68 = x_109; x_69 = x_105; -x_70 = x_106; -x_71 = x_108; -x_72 = x_107; +x_70 = x_108; +x_71 = x_107; +x_72 = x_106; x_73 = x_111; goto block_97; } @@ -24845,19 +24845,19 @@ goto block_97; block_149: { 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 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_114); +x_136 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_117); x_137 = lean_ctor_get(x_136, 0); lean_inc(x_137); lean_dec_ref(x_136); -x_138 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_114, x_131, x_116, x_113, x_130, x_123, x_133, x_119, x_126, x_125); +x_138 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_117, x_121, x_113, x_126, x_114, x_133, x_129, x_131, x_134, x_124); x_139 = lean_ctor_get(x_138, 0); lean_inc(x_139); lean_dec_ref(x_138); -x_140 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_114, x_131, x_116, x_113, x_130, x_123, x_133, x_119, x_126, x_125); +x_140 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_117, x_121, x_113, x_126, x_114, x_133, x_129, x_131, x_134, x_124); x_141 = lean_ctor_get(x_140, 0); lean_inc(x_141); lean_dec_ref(x_140); -x_142 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_134, x_114, x_131, x_116, x_113, x_130, x_123, x_133, x_119, x_126, x_125); +x_142 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_120, x_117, x_121, x_113, x_126, x_114, x_133, x_129, x_131, x_134, x_124); x_143 = lean_ctor_get(x_142, 0); lean_inc(x_143); lean_dec_ref(x_142); @@ -24865,61 +24865,61 @@ lean_inc(x_28); x_144 = l_Lean_mkNatLit(x_28); x_145 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; lean_inc_ref(x_135); -x_146 = l_Lean_mkApp10(x_117, x_132, x_115, x_127, x_120, x_135, x_144, x_145, x_122, x_124, x_129); -x_147 = l_Lean_mkConst(x_128, x_118); +x_146 = l_Lean_mkApp10(x_118, x_116, x_122, x_119, x_125, x_135, x_144, x_145, x_123, x_130, x_132); +x_147 = l_Lean_mkConst(x_127, x_115); x_148 = l_Lean_mkApp5(x_147, x_137, x_139, x_141, x_143, x_135); x_29 = x_146; x_30 = x_148; -x_31 = x_114; -x_32 = x_131; -x_33 = x_116; -x_34 = x_113; -x_35 = x_130; -x_36 = x_123; -x_37 = x_133; -x_38 = x_119; -x_39 = x_126; -x_40 = x_125; +x_31 = x_117; +x_32 = x_121; +x_33 = x_113; +x_34 = x_126; +x_35 = x_114; +x_36 = x_133; +x_37 = x_129; +x_38 = x_131; +x_39 = x_134; +x_40 = x_124; x_41 = lean_box(0); goto block_60; } block_193: { lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_164 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_151); +x_164 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_152); x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); lean_dec_ref(x_164); lean_inc_ref(x_26); -x_166 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +x_166 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_26, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); x_167 = lean_ctor_get(x_166, 0); lean_inc(x_167); lean_dec_ref(x_166); lean_inc_ref(x_27); -x_168 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +x_168 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_27, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); lean_dec_ref(x_168); -x_170 = lean_ctor_get(x_161, 0); +x_170 = lean_ctor_get(x_153, 0); lean_inc(x_170); -x_171 = lean_ctor_get(x_161, 1); +x_171 = lean_ctor_get(x_153, 1); lean_inc_ref(x_171); lean_inc_ref(x_171); -x_172 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_171, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +x_172 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_171, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); x_173 = lean_ctor_get(x_172, 0); lean_inc(x_173); lean_dec_ref(x_172); lean_inc(x_155); -lean_inc_ref(x_156); -lean_inc(x_153); lean_inc_ref(x_162); -lean_inc(x_154); -lean_inc_ref(x_157); -lean_inc(x_150); -lean_inc(x_152); -lean_inc(x_158); +lean_inc(x_160); +lean_inc_ref(x_158); +lean_inc(x_161); lean_inc_ref(x_151); -x_174 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_20, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +lean_inc(x_157); +lean_inc(x_150); +lean_inc(x_154); +lean_inc_ref(x_152); +x_174 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_20, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); if (lean_obj_tag(x_174) == 0) { lean_object* x_175; lean_object* x_176; @@ -24927,16 +24927,16 @@ x_175 = lean_ctor_get(x_174, 0); lean_inc(x_175); lean_dec_ref(x_174); lean_inc(x_155); -lean_inc_ref(x_156); -lean_inc(x_153); lean_inc_ref(x_162); -lean_inc(x_154); -lean_inc_ref(x_157); -lean_inc(x_150); -lean_inc(x_152); -lean_inc(x_158); +lean_inc(x_160); +lean_inc_ref(x_158); +lean_inc(x_161); lean_inc_ref(x_151); -x_176 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_21, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +lean_inc(x_157); +lean_inc(x_150); +lean_inc(x_154); +lean_inc_ref(x_152); +x_176 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_LeCnstr_toExprProof(x_21, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); if (lean_obj_tag(x_176) == 0) { lean_object* x_177; lean_object* x_178; @@ -24944,16 +24944,16 @@ x_177 = lean_ctor_get(x_176, 0); lean_inc(x_177); lean_dec_ref(x_176); lean_inc(x_155); -lean_inc_ref(x_156); -lean_inc(x_153); lean_inc_ref(x_162); -lean_inc(x_154); -lean_inc_ref(x_157); -lean_inc(x_150); -lean_inc(x_152); -lean_inc(x_158); +lean_inc(x_160); +lean_inc_ref(x_158); +lean_inc(x_161); lean_inc_ref(x_151); -x_178 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_161, x_151, x_158, x_152, x_150, x_157, x_154, x_162, x_153, x_156, x_155); +lean_inc(x_157); +lean_inc(x_150); +lean_inc(x_154); +lean_inc_ref(x_152); +x_178 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_153, x_152, x_154, x_150, x_157, x_151, x_161, x_158, x_160, x_162, x_155); if (lean_obj_tag(x_178) == 0) { lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; @@ -24961,7 +24961,7 @@ x_179 = lean_ctor_get(x_178, 0); lean_inc(x_179); lean_dec_ref(x_178); x_180 = lean_box(0); -x_181 = l_Lean_mkConst(x_160, x_180); +x_181 = l_Lean_mkConst(x_156, x_180); x_182 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; x_183 = lean_int_dec_le(x_182, x_170); if (x_183 == 0) @@ -24978,26 +24978,26 @@ x_189 = l_Lean_instToExprInt_mkNat(x_188); x_190 = l_Lean_mkApp3(x_184, x_185, x_186, x_189); x_113 = x_150; x_114 = x_151; -x_115 = x_167; -x_116 = x_152; -x_117 = x_181; -x_118 = x_180; -x_119 = x_153; -x_120 = x_173; -x_121 = lean_box(0); -x_122 = x_175; -x_123 = x_154; -x_124 = x_177; -x_125 = x_155; -x_126 = x_156; -x_127 = x_169; -x_128 = x_163; -x_129 = x_179; -x_130 = x_157; -x_131 = x_158; -x_132 = x_165; -x_133 = x_162; -x_134 = x_171; +x_115 = x_180; +x_116 = x_165; +x_117 = x_152; +x_118 = x_181; +x_119 = x_169; +x_120 = x_171; +x_121 = x_154; +x_122 = x_167; +x_123 = x_175; +x_124 = x_155; +x_125 = x_173; +x_126 = x_157; +x_127 = x_163; +x_128 = lean_box(0); +x_129 = x_158; +x_130 = x_177; +x_131 = x_160; +x_132 = x_179; +x_133 = x_161; +x_134 = x_162; x_135 = x_190; goto block_149; } @@ -25009,26 +25009,26 @@ lean_dec(x_170); x_192 = l_Lean_instToExprInt_mkNat(x_191); x_113 = x_150; x_114 = x_151; -x_115 = x_167; -x_116 = x_152; -x_117 = x_181; -x_118 = x_180; -x_119 = x_153; -x_120 = x_173; -x_121 = lean_box(0); -x_122 = x_175; -x_123 = x_154; -x_124 = x_177; -x_125 = x_155; -x_126 = x_156; -x_127 = x_169; -x_128 = x_163; -x_129 = x_179; -x_130 = x_157; -x_131 = x_158; -x_132 = x_165; -x_133 = x_162; -x_134 = x_171; +x_115 = x_180; +x_116 = x_165; +x_117 = x_152; +x_118 = x_181; +x_119 = x_169; +x_120 = x_171; +x_121 = x_154; +x_122 = x_167; +x_123 = x_175; +x_124 = x_155; +x_125 = x_173; +x_126 = x_157; +x_127 = x_163; +x_128 = lean_box(0); +x_129 = x_158; +x_130 = x_177; +x_131 = x_160; +x_132 = x_179; +x_133 = x_161; +x_134 = x_162; x_135 = x_192; goto block_149; } @@ -25045,14 +25045,14 @@ lean_dec(x_167); lean_dec(x_165); lean_dec(x_163); lean_dec_ref(x_162); +lean_dec(x_161); lean_dec(x_160); -lean_dec(x_158); -lean_dec_ref(x_157); -lean_dec_ref(x_156); +lean_dec_ref(x_158); +lean_dec(x_157); +lean_dec(x_156); lean_dec(x_155); lean_dec(x_154); -lean_dec(x_153); -lean_dec(x_152); +lean_dec_ref(x_152); lean_dec_ref(x_151); lean_dec(x_150); lean_dec(x_28); @@ -25074,15 +25074,15 @@ lean_dec(x_167); lean_dec(x_165); lean_dec(x_163); lean_dec_ref(x_162); -lean_dec_ref(x_161); +lean_dec(x_161); lean_dec(x_160); -lean_dec(x_158); -lean_dec_ref(x_157); -lean_dec_ref(x_156); +lean_dec_ref(x_158); +lean_dec(x_157); +lean_dec(x_156); lean_dec(x_155); lean_dec(x_154); -lean_dec(x_153); -lean_dec(x_152); +lean_dec_ref(x_153); +lean_dec_ref(x_152); lean_dec_ref(x_151); lean_dec(x_150); lean_dec(x_28); @@ -25103,15 +25103,15 @@ lean_dec(x_167); lean_dec(x_165); lean_dec(x_163); lean_dec_ref(x_162); -lean_dec_ref(x_161); +lean_dec(x_161); lean_dec(x_160); -lean_dec(x_158); -lean_dec_ref(x_157); -lean_dec_ref(x_156); +lean_dec_ref(x_158); +lean_dec(x_157); +lean_dec(x_156); lean_dec(x_155); lean_dec(x_154); -lean_dec(x_153); -lean_dec(x_152); +lean_dec_ref(x_153); +lean_dec_ref(x_152); lean_dec_ref(x_151); lean_dec(x_150); lean_dec(x_28); @@ -25129,19 +25129,19 @@ if (x_24 == 0) { lean_object* x_207; x_207 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__8; -x_150 = x_194; +x_150 = x_196; x_151 = x_195; -x_152 = x_196; -x_153 = x_198; -x_154 = x_201; -x_155 = x_203; -x_156 = x_205; +x_152 = x_198; +x_153 = x_201; +x_154 = x_205; +x_155 = x_194; +x_156 = x_206; x_157 = x_197; x_158 = x_199; x_159 = lean_box(0); -x_160 = x_206; -x_161 = x_202; -x_162 = x_204; +x_160 = x_202; +x_161 = x_204; +x_162 = x_203; x_163 = x_207; goto block_193; } @@ -25149,19 +25149,19 @@ else { lean_object* x_208; x_208 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__10; -x_150 = x_194; +x_150 = x_196; x_151 = x_195; -x_152 = x_196; -x_153 = x_198; -x_154 = x_201; -x_155 = x_203; -x_156 = x_205; +x_152 = x_198; +x_153 = x_201; +x_154 = x_205; +x_155 = x_194; +x_156 = x_206; x_157 = x_197; x_158 = x_199; x_159 = lean_box(0); -x_160 = x_206; -x_161 = x_202; -x_162 = x_204; +x_160 = x_202; +x_161 = x_204; +x_162 = x_203; x_163 = x_208; goto block_193; } @@ -25174,16 +25174,16 @@ if (x_24 == 0) { lean_object* x_221; x_221 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__12; -x_98 = x_213; -x_99 = x_210; -x_100 = x_212; -x_101 = x_214; -x_102 = x_217; -x_103 = x_211; +x_98 = x_219; +x_99 = x_212; +x_100 = x_214; +x_101 = x_210; +x_102 = x_213; +x_103 = x_216; x_104 = lean_box(0); -x_105 = x_215; -x_106 = x_219; -x_107 = x_216; +x_105 = x_217; +x_106 = x_211; +x_107 = x_215; x_108 = x_218; x_109 = x_221; goto block_112; @@ -25192,16 +25192,16 @@ else { lean_object* x_222; x_222 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__14; -x_98 = x_213; -x_99 = x_210; -x_100 = x_212; -x_101 = x_214; -x_102 = x_217; -x_103 = x_211; +x_98 = x_219; +x_99 = x_212; +x_100 = x_214; +x_101 = x_210; +x_102 = x_213; +x_103 = x_216; x_104 = lean_box(0); -x_105 = x_215; -x_106 = x_219; -x_107 = x_216; +x_105 = x_217; +x_106 = x_211; +x_107 = x_215; x_108 = x_218; x_109 = x_222; goto block_112; @@ -25216,18 +25216,18 @@ x_223 = lean_ctor_get(x_25, 0); lean_inc(x_223); lean_dec_ref(x_25); x_224 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__16; -x_194 = x_213; -x_195 = x_210; +x_194 = x_219; +x_195 = x_214; x_196 = x_212; -x_197 = x_214; -x_198 = x_217; -x_199 = x_211; +x_197 = x_213; +x_198 = x_210; +x_199 = x_216; x_200 = lean_box(0); -x_201 = x_215; -x_202 = x_223; -x_203 = x_219; -x_204 = x_216; -x_205 = x_218; +x_201 = x_223; +x_202 = x_217; +x_203 = x_218; +x_204 = x_215; +x_205 = x_211; x_206 = x_224; goto block_209; } @@ -25238,18 +25238,18 @@ x_225 = lean_ctor_get(x_25, 0); lean_inc(x_225); lean_dec_ref(x_25); x_226 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof___lam__0___closed__18; -x_194 = x_213; -x_195 = x_210; +x_194 = x_219; +x_195 = x_214; x_196 = x_212; -x_197 = x_214; -x_198 = x_217; -x_199 = x_211; +x_197 = x_213; +x_198 = x_210; +x_199 = x_216; x_200 = lean_box(0); -x_201 = x_215; -x_202 = x_225; -x_203 = x_219; -x_204 = x_216; -x_205 = x_218; +x_201 = x_225; +x_202 = x_217; +x_203 = x_218; +x_204 = x_215; +x_205 = x_211; x_206 = x_226; goto block_209; } @@ -25755,7 +25755,7 @@ return x_1194; { lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_22 = l_Lean_mkApp6(x_17, x_16, x_20, x_19, x_14, x_21, x_18); +x_22 = l_Lean_mkApp6(x_17, x_19, x_20, x_16, x_18, x_21, x_15); x_23 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_23, 0, x_22); return x_23; @@ -25764,7 +25764,7 @@ return x_23; { lean_object* x_32; lean_object* x_33; lean_object* x_34; x_32 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_33 = l_Lean_mkApp6(x_26, x_25, x_31, x_30, x_28, x_32, x_27); +x_33 = l_Lean_mkApp6(x_30, x_29, x_31, x_28, x_26, x_32, x_27); x_34 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_34, 0, x_33); return x_34; @@ -25773,7 +25773,7 @@ return x_34; { lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_46 = l_Lean_mkApp8(x_39, x_43, x_40, x_41, x_42, x_38, x_44, x_45, x_37); +x_46 = l_Lean_mkApp8(x_43, x_39, x_36, x_38, x_42, x_37, x_44, x_45, x_40); x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); return x_47; @@ -25782,36 +25782,36 @@ return x_47; { lean_object* x_60; uint8_t x_61; x_60 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_61 = lean_int_dec_le(x_60, x_50); +x_61 = lean_int_dec_le(x_60, x_56); if (x_61 == 0) { lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; x_62 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_63 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_57); +lean_inc(x_55); x_64 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_57); +lean_ctor_set(x_64, 1, x_55); x_65 = l_Lean_Expr_const___override(x_62, x_64); -lean_inc_ref(x_54); -x_66 = l_Lean_Name_mkStr1(x_54); -lean_inc(x_57); -x_67 = l_Lean_Expr_const___override(x_66, x_57); +lean_inc_ref(x_49); +x_66 = l_Lean_Name_mkStr1(x_49); +lean_inc(x_55); +x_67 = l_Lean_Expr_const___override(x_66, x_55); x_68 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_69 = l_Lean_Name_mkStr2(x_54, x_68); -x_70 = l_Lean_Expr_const___override(x_69, x_57); -x_71 = lean_int_neg(x_50); -lean_dec(x_50); +x_69 = l_Lean_Name_mkStr2(x_49, x_68); +x_70 = l_Lean_Expr_const___override(x_69, x_55); +x_71 = lean_int_neg(x_56); +lean_dec(x_56); x_72 = l_Int_toNat(x_71); lean_dec(x_71); x_73 = l_Lean_instToExprInt_mkNat(x_72); x_74 = l_Lean_mkApp3(x_65, x_67, x_70, x_73); -x_36 = lean_box(0); -x_37 = x_52; -x_38 = x_51; -x_39 = x_53; -x_40 = x_55; -x_41 = x_56; +x_36 = x_50; +x_37 = x_53; +x_38 = x_52; +x_39 = x_51; +x_40 = x_54; +x_41 = lean_box(0); x_42 = x_59; x_43 = x_58; x_44 = x_74; @@ -25820,17 +25820,17 @@ goto block_48; else { lean_object* x_75; lean_object* x_76; -lean_dec(x_57); -lean_dec_ref(x_54); -x_75 = l_Int_toNat(x_50); -lean_dec(x_50); +lean_dec(x_55); +lean_dec_ref(x_49); +x_75 = l_Int_toNat(x_56); +lean_dec(x_56); x_76 = l_Lean_instToExprInt_mkNat(x_75); -x_36 = lean_box(0); -x_37 = x_52; -x_38 = x_51; -x_39 = x_53; -x_40 = x_55; -x_41 = x_56; +x_36 = x_50; +x_37 = x_53; +x_38 = x_52; +x_39 = x_51; +x_40 = x_54; +x_41 = lean_box(0); x_42 = x_59; x_43 = x_58; x_44 = x_76; @@ -25847,36 +25847,36 @@ 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; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; x_91 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_92 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_87); +lean_inc(x_86); x_93 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_87); +lean_ctor_set(x_93, 1, x_86); x_94 = l_Lean_Expr_const___override(x_91, x_93); -lean_inc_ref(x_84); -x_95 = l_Lean_Name_mkStr1(x_84); -lean_inc(x_87); -x_96 = l_Lean_Expr_const___override(x_95, x_87); +lean_inc_ref(x_78); +x_95 = l_Lean_Name_mkStr1(x_78); +lean_inc(x_86); +x_96 = l_Lean_Expr_const___override(x_95, x_86); x_97 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_84); -x_98 = l_Lean_Name_mkStr2(x_84, x_97); -lean_inc(x_87); -x_99 = l_Lean_Expr_const___override(x_98, x_87); +lean_inc_ref(x_78); +x_98 = l_Lean_Name_mkStr2(x_78, x_97); +lean_inc(x_86); +x_99 = l_Lean_Expr_const___override(x_98, x_86); x_100 = lean_int_neg(x_83); lean_dec(x_83); x_101 = l_Int_toNat(x_100); lean_dec(x_100); x_102 = l_Lean_instToExprInt_mkNat(x_101); x_103 = l_Lean_mkApp3(x_94, x_96, x_99, x_102); -x_49 = lean_box(0); -x_50 = x_79; +x_49 = x_78; +x_50 = x_88; x_51 = x_81; x_52 = x_80; -x_53 = x_82; -x_54 = x_84; -x_55 = x_88; +x_53 = x_79; +x_54 = x_82; +x_55 = x_86; x_56 = x_85; -x_57 = x_87; -x_58 = x_86; +x_57 = lean_box(0); +x_58 = x_87; x_59 = x_103; goto block_77; } @@ -25886,16 +25886,16 @@ lean_object* x_104; lean_object* x_105; x_104 = l_Int_toNat(x_83); lean_dec(x_83); x_105 = l_Lean_instToExprInt_mkNat(x_104); -x_49 = lean_box(0); -x_50 = x_79; +x_49 = x_78; +x_50 = x_88; x_51 = x_81; x_52 = x_80; -x_53 = x_82; -x_54 = x_84; -x_55 = x_88; +x_53 = x_79; +x_54 = x_82; +x_55 = x_86; x_56 = x_85; -x_57 = x_87; -x_58 = x_86; +x_57 = lean_box(0); +x_58 = x_87; x_59 = x_105; goto block_77; } @@ -25907,15 +25907,15 @@ x_124 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind x_125 = lean_ctor_get(x_124, 0); lean_inc(x_125); lean_dec_ref(x_124); -x_126 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_114, x_115, x_118, x_122, x_107, x_119, x_121, x_112, x_109, x_116, x_111); +x_126 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_109, x_115, x_117, x_121, x_113, x_111, x_119, x_107, x_120, x_112, x_114); x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); lean_dec_ref(x_126); -x_128 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_113, x_115, x_118, x_122, x_107, x_119, x_121, x_112, x_109, x_116, x_111); +x_128 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_108, x_115, x_117, x_121, x_113, x_111, x_119, x_107, x_120, x_112, x_114); x_129 = lean_ctor_get(x_128, 0); lean_inc(x_129); lean_dec_ref(x_128); -x_130 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_117, x_115, x_118, x_122, x_107, x_119, x_121, x_112, x_109, x_116, x_111); +x_130 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_116, x_115, x_117, x_121, x_113, x_111, x_119, x_107, x_120, x_112, x_114); if (lean_obj_tag(x_130) == 0) { lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; @@ -25926,48 +25926,48 @@ x_132 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind x_133 = lean_box(0); x_134 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__2; x_135 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_136 = lean_int_dec_le(x_135, x_108); +x_136 = lean_int_dec_le(x_135, x_118); if (x_136 == 0) { lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; x_137 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__10; x_138 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__12; x_139 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__15; -x_140 = lean_int_neg(x_108); -lean_dec(x_108); +x_140 = lean_int_neg(x_118); +lean_dec(x_118); x_141 = l_Int_toNat(x_140); lean_dec(x_140); x_142 = l_Lean_instToExprInt_mkNat(x_141); x_143 = l_Lean_mkApp3(x_137, x_138, x_139, x_142); -x_78 = lean_box(0); -x_79 = x_123; -x_80 = x_131; -x_81 = x_129; -x_82 = x_134; -x_83 = x_120; -x_84 = x_132; -x_85 = x_127; -x_86 = x_125; -x_87 = x_133; +x_78 = x_132; +x_79 = x_129; +x_80 = x_127; +x_81 = x_125; +x_82 = x_131; +x_83 = x_110; +x_84 = lean_box(0); +x_85 = x_123; +x_86 = x_133; +x_87 = x_134; x_88 = x_143; goto block_106; } else { lean_object* x_144; lean_object* x_145; -x_144 = l_Int_toNat(x_108); -lean_dec(x_108); +x_144 = l_Int_toNat(x_118); +lean_dec(x_118); x_145 = l_Lean_instToExprInt_mkNat(x_144); -x_78 = lean_box(0); -x_79 = x_123; -x_80 = x_131; -x_81 = x_129; -x_82 = x_134; -x_83 = x_120; -x_84 = x_132; -x_85 = x_127; -x_86 = x_125; -x_87 = x_133; +x_78 = x_132; +x_79 = x_129; +x_80 = x_127; +x_81 = x_125; +x_82 = x_131; +x_83 = x_110; +x_84 = lean_box(0); +x_85 = x_123; +x_86 = x_133; +x_87 = x_134; x_88 = x_145; goto block_106; } @@ -25978,32 +25978,32 @@ lean_dec(x_129); lean_dec(x_127); lean_dec(x_125); lean_dec(x_123); -lean_dec(x_120); -lean_dec(x_108); +lean_dec(x_118); +lean_dec(x_110); return x_130; } } block_184: { lean_object* x_171; -lean_inc(x_151); -lean_inc_ref(x_162); -lean_inc(x_150); -lean_inc_ref(x_153); +lean_inc(x_160); +lean_inc_ref(x_157); lean_inc(x_168); -lean_inc_ref(x_165); -lean_inc(x_147); +lean_inc_ref(x_149); +lean_inc(x_167); +lean_inc_ref(x_153); +lean_inc(x_159); lean_inc(x_169); -lean_inc(x_164); -lean_inc_ref(x_161); -x_171 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_156, x_161, x_164, x_169, x_147, x_165, x_168, x_153, x_150, x_162, x_151); +lean_inc(x_165); +lean_inc_ref(x_163); +x_171 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_152, x_163, x_165, x_169, x_159, x_153, x_167, x_149, x_168, x_157, x_160); if (lean_obj_tag(x_171) == 0) { lean_object* x_172; lean_object* x_173; x_172 = lean_ctor_get(x_171, 0); lean_inc(x_172); lean_dec_ref(x_171); -x_173 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_167, x_161, x_164, x_169, x_147, x_165, x_168, x_153, x_150, x_162, x_151); +x_173 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof(x_161, x_163, x_165, x_169, x_159, x_153, x_167, x_149, x_168, x_157, x_160); if (lean_obj_tag(x_173) == 0) { uint8_t x_174; @@ -26012,7 +26012,7 @@ if (x_174 == 0) { lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; x_175 = lean_ctor_get(x_173, 0); -x_176 = l_Lean_mkApp10(x_148, x_152, x_149, x_158, x_160, x_155, x_163, x_159, x_157, x_166, x_170); +x_176 = l_Lean_mkApp10(x_154, x_162, x_156, x_155, x_151, x_147, x_150, x_166, x_164, x_148, x_170); x_177 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; x_178 = l_Lean_mkApp3(x_176, x_177, x_172, x_175); lean_ctor_set(x_173, 0, x_178); @@ -26024,7 +26024,7 @@ lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; x_179 = lean_ctor_get(x_173, 0); lean_inc(x_179); lean_dec(x_173); -x_180 = l_Lean_mkApp10(x_148, x_152, x_149, x_158, x_160, x_155, x_163, x_159, x_157, x_166, x_170); +x_180 = l_Lean_mkApp10(x_154, x_162, x_156, x_155, x_151, x_147, x_150, x_166, x_164, x_148, x_170); x_181 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; x_182 = l_Lean_mkApp3(x_180, x_181, x_172, x_179); x_183 = lean_alloc_ctor(0, 1, 0); @@ -26037,15 +26037,15 @@ else lean_dec(x_172); lean_dec_ref(x_170); lean_dec_ref(x_166); -lean_dec_ref(x_163); -lean_dec_ref(x_160); -lean_dec_ref(x_159); -lean_dec_ref(x_158); -lean_dec_ref(x_157); +lean_dec_ref(x_164); +lean_dec_ref(x_162); +lean_dec_ref(x_156); lean_dec_ref(x_155); -lean_dec_ref(x_152); -lean_dec_ref(x_149); +lean_dec_ref(x_154); +lean_dec_ref(x_151); +lean_dec_ref(x_150); lean_dec_ref(x_148); +lean_dec_ref(x_147); return x_173; } } @@ -26054,25 +26054,25 @@ else lean_dec_ref(x_170); lean_dec(x_169); lean_dec(x_168); -lean_dec_ref(x_167); +lean_dec(x_167); lean_dec_ref(x_166); -lean_dec_ref(x_165); -lean_dec(x_164); +lean_dec(x_165); +lean_dec_ref(x_164); lean_dec_ref(x_163); lean_dec_ref(x_162); lean_dec_ref(x_161); -lean_dec_ref(x_160); -lean_dec_ref(x_159); -lean_dec_ref(x_158); +lean_dec(x_160); +lean_dec(x_159); lean_dec_ref(x_157); +lean_dec_ref(x_156); lean_dec_ref(x_155); +lean_dec_ref(x_154); lean_dec_ref(x_153); -lean_dec_ref(x_152); -lean_dec(x_151); -lean_dec(x_150); +lean_dec_ref(x_151); +lean_dec_ref(x_150); lean_dec_ref(x_149); lean_dec_ref(x_148); -lean_dec(x_147); +lean_dec_ref(x_147); return x_171; } } @@ -26080,87 +26080,87 @@ return x_171; { lean_object* x_211; uint8_t x_212; x_211 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_212 = lean_int_dec_le(x_211, x_207); +x_212 = lean_int_dec_le(x_211, x_197); if (x_212 == 0) { lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; x_213 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_214 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_192); +lean_inc(x_188); x_215 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_192); +lean_ctor_set(x_215, 1, x_188); x_216 = l_Lean_Expr_const___override(x_213, x_215); -lean_inc_ref(x_190); -x_217 = l_Lean_Name_mkStr1(x_190); -lean_inc(x_192); -x_218 = l_Lean_Expr_const___override(x_217, x_192); +lean_inc_ref(x_198); +x_217 = l_Lean_Name_mkStr1(x_198); +lean_inc(x_188); +x_218 = l_Lean_Expr_const___override(x_217, x_188); x_219 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_220 = l_Lean_Name_mkStr2(x_190, x_219); -x_221 = l_Lean_Expr_const___override(x_220, x_192); -x_222 = lean_int_neg(x_207); -lean_dec(x_207); +x_220 = l_Lean_Name_mkStr2(x_198, x_219); +x_221 = l_Lean_Expr_const___override(x_220, x_188); +x_222 = lean_int_neg(x_197); +lean_dec(x_197); x_223 = l_Int_toNat(x_222); lean_dec(x_222); x_224 = l_Lean_instToExprInt_mkNat(x_223); x_225 = l_Lean_mkApp3(x_216, x_218, x_221, x_224); -x_147 = x_185; -x_148 = x_201; -x_149 = x_186; -x_150 = x_202; -x_151 = x_187; -x_152 = x_189; -x_153 = x_188; -x_154 = lean_box(0); -x_155 = x_191; -x_156 = x_193; -x_157 = x_194; -x_158 = x_204; -x_159 = x_205; -x_160 = x_195; -x_161 = x_206; -x_162 = x_196; -x_163 = x_197; -x_164 = x_208; -x_165 = x_198; -x_166 = x_210; -x_167 = x_199; -x_168 = x_209; -x_169 = x_200; +x_147 = x_192; +x_148 = x_210; +x_149 = x_193; +x_150 = x_185; +x_151 = x_186; +x_152 = x_194; +x_153 = x_195; +x_154 = x_196; +x_155 = x_199; +x_156 = x_201; +x_157 = x_200; +x_158 = lean_box(0); +x_159 = x_204; +x_160 = x_203; +x_161 = x_187; +x_162 = x_205; +x_163 = x_189; +x_164 = x_190; +x_165 = x_206; +x_166 = x_207; +x_167 = x_191; +x_168 = x_208; +x_169 = x_209; x_170 = x_225; goto block_184; } else { lean_object* x_226; lean_object* x_227; -lean_dec(x_192); -lean_dec_ref(x_190); -x_226 = l_Int_toNat(x_207); -lean_dec(x_207); +lean_dec_ref(x_198); +lean_dec(x_188); +x_226 = l_Int_toNat(x_197); +lean_dec(x_197); x_227 = l_Lean_instToExprInt_mkNat(x_226); -x_147 = x_185; -x_148 = x_201; -x_149 = x_186; -x_150 = x_202; -x_151 = x_187; -x_152 = x_189; -x_153 = x_188; -x_154 = lean_box(0); -x_155 = x_191; -x_156 = x_193; -x_157 = x_194; -x_158 = x_204; -x_159 = x_205; -x_160 = x_195; -x_161 = x_206; -x_162 = x_196; -x_163 = x_197; -x_164 = x_208; -x_165 = x_198; -x_166 = x_210; -x_167 = x_199; -x_168 = x_209; -x_169 = x_200; +x_147 = x_192; +x_148 = x_210; +x_149 = x_193; +x_150 = x_185; +x_151 = x_186; +x_152 = x_194; +x_153 = x_195; +x_154 = x_196; +x_155 = x_199; +x_156 = x_201; +x_157 = x_200; +x_158 = lean_box(0); +x_159 = x_204; +x_160 = x_203; +x_161 = x_187; +x_162 = x_205; +x_163 = x_189; +x_164 = x_190; +x_165 = x_206; +x_166 = x_207; +x_167 = x_191; +x_168 = x_208; +x_169 = x_209; x_170 = x_227; goto block_184; } @@ -26169,28 +26169,28 @@ goto block_184; { lean_object* x_255; uint8_t x_256; x_255 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_256 = lean_int_dec_le(x_255, x_245); +x_256 = lean_int_dec_le(x_255, x_239); if (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; x_257 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_258 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_236); +lean_inc(x_232); x_259 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_236); +lean_ctor_set(x_259, 1, x_232); x_260 = l_Lean_Expr_const___override(x_257, x_259); -lean_inc_ref(x_234); -x_261 = l_Lean_Name_mkStr1(x_234); -lean_inc(x_236); -x_262 = l_Lean_Expr_const___override(x_261, x_236); +lean_inc_ref(x_241); +x_261 = l_Lean_Name_mkStr1(x_241); +lean_inc(x_232); +x_262 = l_Lean_Expr_const___override(x_261, x_232); x_263 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_234); -x_264 = l_Lean_Name_mkStr2(x_234, x_263); -lean_inc(x_236); -x_265 = l_Lean_Expr_const___override(x_264, x_236); -x_266 = lean_int_neg(x_245); -lean_dec(x_245); +lean_inc_ref(x_241); +x_264 = l_Lean_Name_mkStr2(x_241, x_263); +lean_inc(x_232); +x_265 = l_Lean_Expr_const___override(x_264, x_232); +x_266 = lean_int_neg(x_239); +lean_dec(x_239); x_267 = l_Int_toNat(x_266); lean_dec(x_266); x_268 = l_Lean_instToExprInt_mkNat(x_267); @@ -26198,22 +26198,22 @@ x_269 = l_Lean_mkApp3(x_260, x_262, x_265, x_268); x_185 = x_229; x_186 = x_230; x_187 = x_231; -x_188 = x_233; -x_189 = x_232; -x_190 = x_234; -x_191 = x_235; -x_192 = x_236; -x_193 = x_237; -x_194 = x_254; -x_195 = x_239; -x_196 = x_238; -x_197 = x_241; -x_198 = x_240; -x_199 = x_242; -x_200 = x_243; -x_201 = x_244; -x_202 = x_246; -x_203 = lean_box(0); +x_188 = x_232; +x_189 = x_233; +x_190 = x_254; +x_191 = x_234; +x_192 = x_235; +x_193 = x_236; +x_194 = x_237; +x_195 = x_238; +x_196 = x_240; +x_197 = x_242; +x_198 = x_241; +x_199 = x_243; +x_200 = x_244; +x_201 = x_245; +x_202 = lean_box(0); +x_203 = x_247; x_204 = x_248; x_205 = x_249; x_206 = x_250; @@ -26226,28 +26226,28 @@ goto block_228; else { lean_object* x_270; lean_object* x_271; -x_270 = l_Int_toNat(x_245); -lean_dec(x_245); +x_270 = l_Int_toNat(x_239); +lean_dec(x_239); x_271 = l_Lean_instToExprInt_mkNat(x_270); x_185 = x_229; x_186 = x_230; x_187 = x_231; -x_188 = x_233; -x_189 = x_232; -x_190 = x_234; -x_191 = x_235; -x_192 = x_236; -x_193 = x_237; -x_194 = x_254; -x_195 = x_239; -x_196 = x_238; -x_197 = x_241; -x_198 = x_240; -x_199 = x_242; -x_200 = x_243; -x_201 = x_244; -x_202 = x_246; -x_203 = lean_box(0); +x_188 = x_232; +x_189 = x_233; +x_190 = x_254; +x_191 = x_234; +x_192 = x_235; +x_193 = x_236; +x_194 = x_237; +x_195 = x_238; +x_196 = x_240; +x_197 = x_242; +x_198 = x_241; +x_199 = x_243; +x_200 = x_244; +x_201 = x_245; +x_202 = lean_box(0); +x_203 = x_247; x_204 = x_248; x_205 = x_249; x_206 = x_250; @@ -26262,51 +26262,51 @@ goto block_228; { lean_object* x_299; uint8_t x_300; x_299 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_300 = lean_int_dec_le(x_299, x_285); +x_300 = lean_int_dec_le(x_299, x_278); if (x_300 == 0) { lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; x_301 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_302 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_280); +lean_inc(x_275); x_303 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_303, 0, x_302); -lean_ctor_set(x_303, 1, x_280); +lean_ctor_set(x_303, 1, x_275); x_304 = l_Lean_Expr_const___override(x_301, x_303); -lean_inc_ref(x_278); -x_305 = l_Lean_Name_mkStr1(x_278); -lean_inc(x_280); -x_306 = l_Lean_Expr_const___override(x_305, x_280); +lean_inc_ref(x_286); +x_305 = l_Lean_Name_mkStr1(x_286); +lean_inc(x_275); +x_306 = l_Lean_Expr_const___override(x_305, x_275); x_307 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_278); -x_308 = l_Lean_Name_mkStr2(x_278, x_307); -lean_inc(x_280); -x_309 = l_Lean_Expr_const___override(x_308, x_280); -x_310 = lean_int_neg(x_285); -lean_dec(x_285); +lean_inc_ref(x_286); +x_308 = l_Lean_Name_mkStr2(x_286, x_307); +lean_inc(x_275); +x_309 = l_Lean_Expr_const___override(x_308, x_275); +x_310 = lean_int_neg(x_278); +lean_dec(x_278); x_311 = l_Int_toNat(x_310); lean_dec(x_310); x_312 = l_Lean_instToExprInt_mkNat(x_311); x_313 = l_Lean_mkApp3(x_304, x_306, x_309, x_312); -x_229 = x_273; -x_230 = x_274; -x_231 = x_275; -x_232 = x_277; +x_229 = x_298; +x_230 = x_273; +x_231 = x_274; +x_232 = x_275; x_233 = x_276; -x_234 = x_278; +x_234 = x_277; x_235 = x_279; x_236 = x_280; x_237 = x_281; -x_238 = x_283; -x_239 = x_282; +x_238 = x_282; +x_239 = x_283; x_240 = x_284; -x_241 = x_298; -x_242 = x_286; +x_241 = x_286; +x_242 = x_285; x_243 = x_287; x_244 = x_288; x_245 = x_289; -x_246 = x_290; -x_247 = lean_box(0); +x_246 = lean_box(0); +x_247 = x_291; x_248 = x_292; x_249 = x_293; x_250 = x_294; @@ -26319,28 +26319,28 @@ goto block_272; else { lean_object* x_314; lean_object* x_315; -x_314 = l_Int_toNat(x_285); -lean_dec(x_285); +x_314 = l_Int_toNat(x_278); +lean_dec(x_278); x_315 = l_Lean_instToExprInt_mkNat(x_314); -x_229 = x_273; -x_230 = x_274; -x_231 = x_275; -x_232 = x_277; +x_229 = x_298; +x_230 = x_273; +x_231 = x_274; +x_232 = x_275; x_233 = x_276; -x_234 = x_278; +x_234 = x_277; x_235 = x_279; x_236 = x_280; x_237 = x_281; -x_238 = x_283; -x_239 = x_282; +x_238 = x_282; +x_239 = x_283; x_240 = x_284; -x_241 = x_298; -x_242 = x_286; +x_241 = x_286; +x_242 = x_285; x_243 = x_287; x_244 = x_288; x_245 = x_289; -x_246 = x_290; -x_247 = lean_box(0); +x_246 = lean_box(0); +x_247 = x_291; x_248 = x_292; x_249 = x_293; x_250 = x_294; @@ -26355,51 +26355,51 @@ goto block_272; { lean_object* x_343; uint8_t x_344; x_343 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_344 = lean_int_dec_le(x_343, x_330); +x_344 = lean_int_dec_le(x_343, x_317); if (x_344 == 0) { lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; x_345 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_346 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_324); +lean_inc(x_319); x_347 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_347, 0, x_346); -lean_ctor_set(x_347, 1, x_324); +lean_ctor_set(x_347, 1, x_319); x_348 = l_Lean_Expr_const___override(x_345, x_347); -lean_inc_ref(x_322); -x_349 = l_Lean_Name_mkStr1(x_322); -lean_inc(x_324); -x_350 = l_Lean_Expr_const___override(x_349, x_324); +lean_inc_ref(x_329); +x_349 = l_Lean_Name_mkStr1(x_329); +lean_inc(x_319); +x_350 = l_Lean_Expr_const___override(x_349, x_319); x_351 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_322); -x_352 = l_Lean_Name_mkStr2(x_322, x_351); -lean_inc(x_324); -x_353 = l_Lean_Expr_const___override(x_352, x_324); -x_354 = lean_int_neg(x_330); -lean_dec(x_330); +lean_inc_ref(x_329); +x_352 = l_Lean_Name_mkStr2(x_329, x_351); +lean_inc(x_319); +x_353 = l_Lean_Expr_const___override(x_352, x_319); +x_354 = lean_int_neg(x_317); +lean_dec(x_317); x_355 = l_Int_toNat(x_354); lean_dec(x_354); x_356 = l_Lean_instToExprInt_mkNat(x_355); x_357 = l_Lean_mkApp3(x_348, x_350, x_353, x_356); -x_273 = x_317; +x_273 = x_342; x_274 = x_318; x_275 = x_319; -x_276 = x_321; -x_277 = x_320; +x_276 = x_320; +x_277 = x_321; x_278 = x_322; x_279 = x_323; x_280 = x_324; x_281 = x_325; -x_282 = x_342; -x_283 = x_326; -x_284 = x_327; -x_285 = x_328; +x_282 = x_326; +x_283 = x_327; +x_284 = x_328; +x_285 = x_330; x_286 = x_329; x_287 = x_331; x_288 = x_332; x_289 = x_333; -x_290 = x_334; -x_291 = lean_box(0); +x_290 = lean_box(0); +x_291 = x_335; x_292 = x_336; x_293 = x_337; x_294 = x_338; @@ -26412,28 +26412,28 @@ goto block_316; else { lean_object* x_358; lean_object* x_359; -x_358 = l_Int_toNat(x_330); -lean_dec(x_330); +x_358 = l_Int_toNat(x_317); +lean_dec(x_317); x_359 = l_Lean_instToExprInt_mkNat(x_358); -x_273 = x_317; +x_273 = x_342; x_274 = x_318; x_275 = x_319; -x_276 = x_321; -x_277 = x_320; +x_276 = x_320; +x_277 = x_321; x_278 = x_322; x_279 = x_323; x_280 = x_324; x_281 = x_325; -x_282 = x_342; -x_283 = x_326; -x_284 = x_327; -x_285 = x_328; +x_282 = x_326; +x_283 = x_327; +x_284 = x_328; +x_285 = x_330; x_286 = x_329; x_287 = x_331; x_288 = x_332; x_289 = x_333; -x_290 = x_334; -x_291 = lean_box(0); +x_290 = lean_box(0); +x_291 = x_335; x_292 = x_336; x_293 = x_337; x_294 = x_338; @@ -26448,51 +26448,51 @@ goto block_316; { lean_object* x_387; uint8_t x_388; x_387 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_388 = lean_int_dec_le(x_387, x_362); +x_388 = lean_int_dec_le(x_387, x_371); if (x_388 == 0) { lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; x_389 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_390 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_368); +lean_inc(x_363); x_391 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_391, 0, x_390); -lean_ctor_set(x_391, 1, x_368); +lean_ctor_set(x_391, 1, x_363); x_392 = l_Lean_Expr_const___override(x_389, x_391); -lean_inc_ref(x_366); -x_393 = l_Lean_Name_mkStr1(x_366); -lean_inc(x_368); -x_394 = l_Lean_Expr_const___override(x_393, x_368); +lean_inc_ref(x_374); +x_393 = l_Lean_Name_mkStr1(x_374); +lean_inc(x_363); +x_394 = l_Lean_Expr_const___override(x_393, x_363); x_395 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_366); -x_396 = l_Lean_Name_mkStr2(x_366, x_395); -lean_inc(x_368); -x_397 = l_Lean_Expr_const___override(x_396, x_368); -x_398 = lean_int_neg(x_362); -lean_dec(x_362); +lean_inc_ref(x_374); +x_396 = l_Lean_Name_mkStr2(x_374, x_395); +lean_inc(x_363); +x_397 = l_Lean_Expr_const___override(x_396, x_363); +x_398 = lean_int_neg(x_371); +lean_dec(x_371); x_399 = l_Int_toNat(x_398); lean_dec(x_398); x_400 = l_Lean_instToExprInt_mkNat(x_399); x_401 = l_Lean_mkApp3(x_392, x_394, x_397, x_400); x_317 = x_361; -x_318 = x_386; +x_318 = x_362; x_319 = x_363; -x_320 = x_365; -x_321 = x_364; +x_320 = x_364; +x_321 = x_365; x_322 = x_366; x_323 = x_367; x_324 = x_368; x_325 = x_369; x_326 = x_370; -x_327 = x_371; -x_328 = x_372; -x_329 = x_373; -x_330 = x_374; -x_331 = x_375; -x_332 = x_376; -x_333 = x_377; -x_334 = x_378; -x_335 = lean_box(0); +x_327 = x_372; +x_328 = x_373; +x_329 = x_374; +x_330 = x_375; +x_331 = x_376; +x_332 = x_377; +x_333 = x_386; +x_334 = lean_box(0); +x_335 = x_379; x_336 = x_380; x_337 = x_381; x_338 = x_382; @@ -26505,28 +26505,28 @@ goto block_360; else { lean_object* x_402; lean_object* x_403; -x_402 = l_Int_toNat(x_362); -lean_dec(x_362); +x_402 = l_Int_toNat(x_371); +lean_dec(x_371); x_403 = l_Lean_instToExprInt_mkNat(x_402); x_317 = x_361; -x_318 = x_386; +x_318 = x_362; x_319 = x_363; -x_320 = x_365; -x_321 = x_364; +x_320 = x_364; +x_321 = x_365; x_322 = x_366; x_323 = x_367; x_324 = x_368; x_325 = x_369; x_326 = x_370; -x_327 = x_371; -x_328 = x_372; -x_329 = x_373; -x_330 = x_374; -x_331 = x_375; -x_332 = x_376; -x_333 = x_377; -x_334 = x_378; -x_335 = lean_box(0); +x_327 = x_372; +x_328 = x_373; +x_329 = x_374; +x_330 = x_375; +x_331 = x_376; +x_332 = x_377; +x_333 = x_386; +x_334 = lean_box(0); +x_335 = x_379; x_336 = x_380; x_337 = x_381; x_338 = x_382; @@ -26541,7 +26541,7 @@ goto block_360; { lean_object* x_416; lean_object* x_417; lean_object* x_418; x_416 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_417 = l_Lean_mkApp10(x_407, x_406, x_413, x_409, x_414, x_410, x_415, x_405, x_416, x_412, x_408); +x_417 = l_Lean_mkApp10(x_413, x_408, x_411, x_409, x_407, x_406, x_415, x_405, x_416, x_412, x_410); x_418 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_418, 0, x_417); return x_418; @@ -26550,61 +26550,61 @@ return x_418; { lean_object* x_433; uint8_t x_434; x_433 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_434 = lean_int_dec_le(x_433, x_427); +x_434 = lean_int_dec_le(x_433, x_422); if (x_434 == 0) { lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; x_435 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_436 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_422); +lean_inc(x_424); x_437 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_437, 0, x_436); -lean_ctor_set(x_437, 1, x_422); +lean_ctor_set(x_437, 1, x_424); x_438 = l_Lean_Expr_const___override(x_435, x_437); -lean_inc_ref(x_429); -x_439 = l_Lean_Name_mkStr1(x_429); -lean_inc(x_422); -x_440 = l_Lean_Expr_const___override(x_439, x_422); +lean_inc_ref(x_430); +x_439 = l_Lean_Name_mkStr1(x_430); +lean_inc(x_424); +x_440 = l_Lean_Expr_const___override(x_439, x_424); x_441 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_442 = l_Lean_Name_mkStr2(x_429, x_441); -x_443 = l_Lean_Expr_const___override(x_442, x_422); -x_444 = lean_int_neg(x_427); -lean_dec(x_427); +x_442 = l_Lean_Name_mkStr2(x_430, x_441); +x_443 = l_Lean_Expr_const___override(x_442, x_424); +x_444 = lean_int_neg(x_422); +lean_dec(x_422); x_445 = l_Int_toNat(x_444); lean_dec(x_444); x_446 = l_Lean_instToExprInt_mkNat(x_445); x_447 = l_Lean_mkApp3(x_438, x_440, x_443, x_446); x_405 = x_420; x_406 = x_421; -x_407 = x_423; -x_408 = x_425; -x_409 = x_424; -x_410 = x_426; -x_411 = lean_box(0); -x_412 = x_430; -x_413 = x_431; -x_414 = x_432; +x_407 = x_432; +x_408 = x_423; +x_409 = x_425; +x_410 = x_427; +x_411 = x_426; +x_412 = x_429; +x_413 = x_428; +x_414 = lean_box(0); x_415 = x_447; goto block_419; } else { lean_object* x_448; lean_object* x_449; -lean_dec_ref(x_429); +lean_dec_ref(x_430); +lean_dec(x_424); +x_448 = l_Int_toNat(x_422); lean_dec(x_422); -x_448 = l_Int_toNat(x_427); -lean_dec(x_427); x_449 = l_Lean_instToExprInt_mkNat(x_448); x_405 = x_420; x_406 = x_421; -x_407 = x_423; -x_408 = x_425; -x_409 = x_424; -x_410 = x_426; -x_411 = lean_box(0); -x_412 = x_430; -x_413 = x_431; -x_414 = x_432; +x_407 = x_432; +x_408 = x_423; +x_409 = x_425; +x_410 = x_427; +x_411 = x_426; +x_412 = x_429; +x_413 = x_428; +x_414 = lean_box(0); x_415 = x_449; goto block_419; } @@ -26613,65 +26613,65 @@ goto block_419; { lean_object* x_464; uint8_t x_465; x_464 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_465 = lean_int_dec_le(x_464, x_451); +x_465 = lean_int_dec_le(x_464, x_454); if (x_465 == 0) { lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; x_466 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_467 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_454); +lean_inc(x_457); x_468 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_468, 0, x_467); -lean_ctor_set(x_468, 1, x_454); +lean_ctor_set(x_468, 1, x_457); x_469 = l_Lean_Expr_const___override(x_466, x_468); -lean_inc_ref(x_462); -x_470 = l_Lean_Name_mkStr1(x_462); -lean_inc(x_454); -x_471 = l_Lean_Expr_const___override(x_470, x_454); +lean_inc_ref(x_461); +x_470 = l_Lean_Name_mkStr1(x_461); +lean_inc(x_457); +x_471 = l_Lean_Expr_const___override(x_470, x_457); x_472 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -lean_inc_ref(x_462); -x_473 = l_Lean_Name_mkStr2(x_462, x_472); -lean_inc(x_454); -x_474 = l_Lean_Expr_const___override(x_473, x_454); -x_475 = lean_int_neg(x_451); -lean_dec(x_451); +lean_inc_ref(x_461); +x_473 = l_Lean_Name_mkStr2(x_461, x_472); +lean_inc(x_457); +x_474 = l_Lean_Expr_const___override(x_473, x_457); +x_475 = lean_int_neg(x_454); +lean_dec(x_454); x_476 = l_Int_toNat(x_475); lean_dec(x_475); x_477 = l_Lean_instToExprInt_mkNat(x_476); x_478 = l_Lean_mkApp3(x_469, x_471, x_474, x_477); -x_420 = x_452; -x_421 = x_453; -x_422 = x_454; +x_420 = x_451; +x_421 = x_452; +x_422 = x_453; x_423 = x_455; x_424 = x_457; x_425 = x_456; -x_426 = x_458; -x_427 = x_459; -x_428 = lean_box(0); -x_429 = x_462; +x_426 = x_463; +x_427 = x_458; +x_428 = x_460; +x_429 = x_459; x_430 = x_461; -x_431 = x_463; +x_431 = lean_box(0); x_432 = x_478; goto block_450; } else { lean_object* x_479; lean_object* x_480; -x_479 = l_Int_toNat(x_451); -lean_dec(x_451); +x_479 = l_Int_toNat(x_454); +lean_dec(x_454); x_480 = l_Lean_instToExprInt_mkNat(x_479); -x_420 = x_452; -x_421 = x_453; -x_422 = x_454; +x_420 = x_451; +x_421 = x_452; +x_422 = x_453; x_423 = x_455; x_424 = x_457; x_425 = x_456; -x_426 = x_458; -x_427 = x_459; -x_428 = lean_box(0); -x_429 = x_462; +x_426 = x_463; +x_427 = x_458; +x_428 = x_460; +x_429 = x_459; x_430 = x_461; -x_431 = x_463; +x_431 = lean_box(0); x_432 = x_480; goto block_450; } @@ -26680,7 +26680,7 @@ goto block_450; { lean_object* x_490; lean_object* x_491; lean_object* x_492; x_490 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_491 = l_Lean_mkApp7(x_484, x_482, x_488, x_486, x_489, x_487, x_490, x_485); +x_491 = l_Lean_mkApp7(x_488, x_482, x_483, x_485, x_489, x_484, x_490, x_486); x_492 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_492, 0, x_491); return x_492; @@ -26689,55 +26689,55 @@ return x_492; { lean_object* x_504; uint8_t x_505; x_504 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_505 = lean_int_dec_le(x_504, x_500); +x_505 = lean_int_dec_le(x_504, x_499); if (x_505 == 0) { lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; x_506 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_507 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_494); +lean_inc(x_496); x_508 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_508, 0, x_507); -lean_ctor_set(x_508, 1, x_494); +lean_ctor_set(x_508, 1, x_496); x_509 = l_Lean_Expr_const___override(x_506, x_508); -lean_inc_ref(x_501); -x_510 = l_Lean_Name_mkStr1(x_501); -lean_inc(x_494); -x_511 = l_Lean_Expr_const___override(x_510, x_494); +lean_inc_ref(x_494); +x_510 = l_Lean_Name_mkStr1(x_494); +lean_inc(x_496); +x_511 = l_Lean_Expr_const___override(x_510, x_496); x_512 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_513 = l_Lean_Name_mkStr2(x_501, x_512); -x_514 = l_Lean_Expr_const___override(x_513, x_494); -x_515 = lean_int_neg(x_500); -lean_dec(x_500); +x_513 = l_Lean_Name_mkStr2(x_494, x_512); +x_514 = l_Lean_Expr_const___override(x_513, x_496); +x_515 = lean_int_neg(x_499); +lean_dec(x_499); x_516 = l_Int_toNat(x_515); lean_dec(x_515); x_517 = l_Lean_instToExprInt_mkNat(x_516); x_518 = l_Lean_mkApp3(x_509, x_511, x_514, x_517); x_482 = x_495; -x_483 = lean_box(0); -x_484 = x_497; -x_485 = x_499; -x_486 = x_498; -x_487 = x_502; -x_488 = x_503; +x_483 = x_503; +x_484 = x_498; +x_485 = x_497; +x_486 = x_500; +x_487 = lean_box(0); +x_488 = x_502; x_489 = x_518; goto block_493; } else { lean_object* x_519; lean_object* x_520; -lean_dec_ref(x_501); -lean_dec(x_494); -x_519 = l_Int_toNat(x_500); -lean_dec(x_500); +lean_dec(x_496); +lean_dec_ref(x_494); +x_519 = l_Int_toNat(x_499); +lean_dec(x_499); x_520 = l_Lean_instToExprInt_mkNat(x_519); x_482 = x_495; -x_483 = lean_box(0); -x_484 = x_497; -x_485 = x_499; -x_486 = x_498; -x_487 = x_502; -x_488 = x_503; +x_483 = x_503; +x_484 = x_498; +x_485 = x_497; +x_486 = x_500; +x_487 = lean_box(0); +x_488 = x_502; x_489 = x_520; goto block_493; } @@ -26746,7 +26746,7 @@ goto block_493; { lean_object* x_530; lean_object* x_531; lean_object* x_532; x_530 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_531 = l_Lean_mkApp7(x_528, x_524, x_523, x_522, x_529, x_527, x_530, x_526); +x_531 = l_Lean_mkApp7(x_524, x_522, x_523, x_526, x_529, x_527, x_530, x_525); x_532 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_532, 0, x_531); return x_532; @@ -26755,7 +26755,7 @@ return x_532; { lean_object* x_545; lean_object* x_546; lean_object* x_547; x_545 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_546 = l_Lean_mkApp10(x_542, x_541, x_538, x_534, x_535, x_540, x_544, x_537, x_545, x_536, x_543); +x_546 = l_Lean_mkApp10(x_537, x_542, x_541, x_534, x_538, x_540, x_544, x_536, x_545, x_543, x_539); x_547 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_547, 0, x_546); return x_547; @@ -26764,39 +26764,39 @@ return x_547; { lean_object* x_562; uint8_t x_563; x_562 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_563 = lean_int_dec_le(x_562, x_554); +x_563 = lean_int_dec_le(x_562, x_553); if (x_563 == 0) { lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; x_564 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_565 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_552); +lean_inc(x_551); x_566 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_566, 0, x_565); -lean_ctor_set(x_566, 1, x_552); +lean_ctor_set(x_566, 1, x_551); x_567 = l_Lean_Expr_const___override(x_564, x_566); -lean_inc_ref(x_558); -x_568 = l_Lean_Name_mkStr1(x_558); -lean_inc(x_552); -x_569 = l_Lean_Expr_const___override(x_568, x_552); +lean_inc_ref(x_555); +x_568 = l_Lean_Name_mkStr1(x_555); +lean_inc(x_551); +x_569 = l_Lean_Expr_const___override(x_568, x_551); x_570 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_571 = l_Lean_Name_mkStr2(x_558, x_570); -x_572 = l_Lean_Expr_const___override(x_571, x_552); -x_573 = lean_int_neg(x_554); -lean_dec(x_554); +x_571 = l_Lean_Name_mkStr2(x_555, x_570); +x_572 = l_Lean_Expr_const___override(x_571, x_551); +x_573 = lean_int_neg(x_553); +lean_dec(x_553); x_574 = l_Int_toNat(x_573); lean_dec(x_573); x_575 = l_Lean_instToExprInt_mkNat(x_574); x_576 = l_Lean_mkApp3(x_567, x_569, x_572, x_575); x_534 = x_549; -x_535 = x_561; -x_536 = x_550; -x_537 = x_551; -x_538 = x_553; -x_539 = lean_box(0); +x_535 = lean_box(0); +x_536 = x_552; +x_537 = x_554; +x_538 = x_561; +x_539 = x_557; x_540 = x_556; -x_541 = x_555; -x_542 = x_559; +x_541 = x_559; +x_542 = x_558; x_543 = x_560; x_544 = x_576; goto block_548; @@ -26804,20 +26804,20 @@ goto block_548; else { lean_object* x_577; lean_object* x_578; -lean_dec_ref(x_558); -lean_dec(x_552); -x_577 = l_Int_toNat(x_554); -lean_dec(x_554); +lean_dec_ref(x_555); +lean_dec(x_551); +x_577 = l_Int_toNat(x_553); +lean_dec(x_553); x_578 = l_Lean_instToExprInt_mkNat(x_577); x_534 = x_549; -x_535 = x_561; -x_536 = x_550; -x_537 = x_551; -x_538 = x_553; -x_539 = lean_box(0); +x_535 = lean_box(0); +x_536 = x_552; +x_537 = x_554; +x_538 = x_561; +x_539 = x_557; x_540 = x_556; -x_541 = x_555; -x_542 = x_559; +x_541 = x_559; +x_542 = x_558; x_543 = x_560; x_544 = x_578; goto block_548; @@ -26827,7 +26827,7 @@ goto block_548; { lean_object* x_589; lean_object* x_590; lean_object* x_591; x_589 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_590 = l_Lean_mkApp8(x_584, x_582, x_585, x_583, x_580, x_588, x_586, x_581, x_589); +x_590 = l_Lean_mkApp8(x_582, x_587, x_580, x_581, x_586, x_588, x_584, x_583, x_589); x_591 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_591, 0, x_590); return x_591; @@ -26835,23 +26835,23 @@ return x_591; block_635: { lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; -x_611 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_601); +x_611 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_604); x_612 = lean_ctor_get(x_611, 0); lean_inc(x_612); lean_dec_ref(x_611); -x_613 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_606, x_601, x_603, x_609, x_593, x_604, x_608, x_600, x_594, x_602, x_596); +x_613 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_600, x_604, x_605, x_608, x_601, x_598, x_606, x_593, x_607, x_599, x_602); x_614 = lean_ctor_get(x_613, 0); lean_inc(x_614); lean_dec_ref(x_613); -x_615 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_597, x_601, x_603, x_609, x_593, x_604, x_608, x_600, x_594, x_602, x_596); +x_615 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_594, x_604, x_605, x_608, x_601, x_598, x_606, x_593, x_607, x_599, x_602); x_616 = lean_ctor_get(x_615, 0); lean_inc(x_616); lean_dec_ref(x_615); -x_617 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_599, x_601, x_603, x_609, x_593, x_604, x_608, x_600, x_594, x_602, x_596); +x_617 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_596, x_604, x_605, x_608, x_601, x_598, x_606, x_593, x_607, x_599, x_602); x_618 = lean_ctor_get(x_617, 0); lean_inc(x_618); lean_dec_ref(x_617); -x_619 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_598, x_601, x_603, x_609, x_593, x_604, x_608, x_600, x_594, x_602, x_596); +x_619 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_603, x_604, x_605, x_608, x_601, x_598, x_606, x_593, x_607, x_599, x_602); if (lean_obj_tag(x_619) == 0) { lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; uint8_t x_625; @@ -26860,46 +26860,46 @@ lean_inc(x_620); lean_dec_ref(x_619); x_621 = lean_box(0); x_622 = l_Lean_mkConst(x_610, x_621); -x_623 = l_Lean_mkNatLit(x_605); +x_623 = l_Lean_mkNatLit(x_595); x_624 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_625 = lean_int_dec_le(x_624, x_607); +x_625 = lean_int_dec_le(x_624, x_597); if (x_625 == 0) { lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; x_626 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__10; x_627 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__12; x_628 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__15; -x_629 = lean_int_neg(x_607); -lean_dec(x_607); +x_629 = lean_int_neg(x_597); +lean_dec(x_597); x_630 = l_Int_toNat(x_629); lean_dec(x_629); x_631 = l_Lean_instToExprInt_mkNat(x_630); x_632 = l_Lean_mkApp3(x_626, x_627, x_628, x_631); -x_580 = x_623; -x_581 = x_620; -x_582 = x_612; -x_583 = x_616; -x_584 = x_622; -x_585 = x_614; -x_586 = x_618; -x_587 = lean_box(0); +x_580 = x_614; +x_581 = x_616; +x_582 = x_622; +x_583 = x_620; +x_584 = x_618; +x_585 = lean_box(0); +x_586 = x_623; +x_587 = x_612; x_588 = x_632; goto block_592; } else { lean_object* x_633; lean_object* x_634; -x_633 = l_Int_toNat(x_607); -lean_dec(x_607); +x_633 = l_Int_toNat(x_597); +lean_dec(x_597); x_634 = l_Lean_instToExprInt_mkNat(x_633); -x_580 = x_623; -x_581 = x_620; -x_582 = x_612; -x_583 = x_616; -x_584 = x_622; -x_585 = x_614; -x_586 = x_618; -x_587 = lean_box(0); +x_580 = x_614; +x_581 = x_616; +x_582 = x_622; +x_583 = x_620; +x_584 = x_618; +x_585 = lean_box(0); +x_586 = x_623; +x_587 = x_612; x_588 = x_634; goto block_592; } @@ -26911,8 +26911,8 @@ lean_dec(x_616); lean_dec(x_614); lean_dec(x_612); lean_dec(x_610); -lean_dec(x_607); -lean_dec(x_605); +lean_dec(x_597); +lean_dec(x_595); return x_619; } } @@ -26920,7 +26920,7 @@ return x_619; { lean_object* x_647; lean_object* x_648; lean_object* x_649; x_647 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_648 = l_Lean_mkApp10(x_638, x_639, x_643, x_645, x_642, x_644, x_640, x_646, x_637, x_641, x_647); +x_648 = l_Lean_mkApp10(x_637, x_642, x_645, x_643, x_641, x_638, x_639, x_646, x_644, x_636, x_647); x_649 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_649, 0, x_648); return x_649; @@ -26928,9 +26928,9 @@ return x_649; block_680: { lean_object* x_663; lean_object* x_664; uint8_t x_665; -x_663 = l_Lean_mkNatLit(x_654); +x_663 = l_Lean_mkNatLit(x_653); x_664 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_665 = lean_int_dec_le(x_664, x_658); +x_665 = lean_int_dec_le(x_664, x_654); if (x_665 == 0) { lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; @@ -26946,22 +26946,22 @@ lean_inc(x_656); x_671 = l_Lean_Expr_const___override(x_670, x_656); x_672 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__14; x_673 = l_Lean_Expr_const___override(x_672, x_656); -x_674 = lean_int_neg(x_658); -lean_dec(x_658); +x_674 = lean_int_neg(x_654); +lean_dec(x_654); x_675 = l_Int_toNat(x_674); lean_dec(x_674); x_676 = l_Lean_instToExprInt_mkNat(x_675); x_677 = l_Lean_mkApp3(x_669, x_671, x_673, x_676); -x_636 = lean_box(0); -x_637 = x_652; -x_638 = x_651; -x_639 = x_655; -x_640 = x_663; -x_641 = x_657; -x_642 = x_659; -x_643 = x_660; -x_644 = x_662; -x_645 = x_661; +x_636 = x_652; +x_637 = x_651; +x_638 = x_662; +x_639 = x_663; +x_640 = lean_box(0); +x_641 = x_658; +x_642 = x_657; +x_643 = x_659; +x_644 = x_661; +x_645 = x_660; x_646 = x_677; goto block_650; } @@ -26969,19 +26969,19 @@ else { lean_object* x_678; lean_object* x_679; lean_dec(x_656); -x_678 = l_Int_toNat(x_658); -lean_dec(x_658); +x_678 = l_Int_toNat(x_654); +lean_dec(x_654); x_679 = l_Lean_instToExprInt_mkNat(x_678); -x_636 = lean_box(0); -x_637 = x_652; -x_638 = x_651; -x_639 = x_655; -x_640 = x_663; -x_641 = x_657; -x_642 = x_659; -x_643 = x_660; -x_644 = x_662; -x_645 = x_661; +x_636 = x_652; +x_637 = x_651; +x_638 = x_662; +x_639 = x_663; +x_640 = lean_box(0); +x_641 = x_658; +x_642 = x_657; +x_643 = x_659; +x_644 = x_661; +x_645 = x_660; x_646 = x_679; goto block_650; } @@ -26989,32 +26989,32 @@ goto block_650; block_727: { lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; -x_700 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_690); +x_700 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_693); x_701 = lean_ctor_get(x_700, 0); lean_inc(x_701); lean_dec_ref(x_700); -x_702 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_695, x_690, x_692, x_698, x_681, x_693, x_697, x_688, x_682, x_691, x_684); +x_702 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_689, x_693, x_694, x_697, x_690, x_686, x_695, x_681, x_696, x_688, x_691); x_703 = lean_ctor_get(x_702, 0); lean_inc(x_703); lean_dec_ref(x_702); -x_704 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_685, x_690, x_692, x_698, x_681, x_693, x_697, x_688, x_682, x_691, x_684); +x_704 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_682, x_693, x_694, x_697, x_690, x_686, x_695, x_681, x_696, x_688, x_691); x_705 = lean_ctor_get(x_704, 0); lean_inc(x_705); lean_dec_ref(x_704); -x_706 = lean_ctor_get(x_689, 0); +x_706 = lean_ctor_get(x_687, 0); lean_inc(x_706); -x_707 = lean_ctor_get(x_689, 1); +x_707 = lean_ctor_get(x_687, 1); lean_inc_ref(x_707); -lean_dec_ref(x_689); -x_708 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_707, x_690, x_692, x_698, x_681, x_693, x_697, x_688, x_682, x_691, x_684); +lean_dec_ref(x_687); +x_708 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_707, x_693, x_694, x_697, x_690, x_686, x_695, x_681, x_696, x_688, x_691); x_709 = lean_ctor_get(x_708, 0); lean_inc(x_709); lean_dec_ref(x_708); -x_710 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_687, x_690, x_692, x_698, x_681, x_693, x_697, x_688, x_682, x_691, x_684); +x_710 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_684, x_693, x_694, x_697, x_690, x_686, x_695, x_681, x_696, x_688, x_691); x_711 = lean_ctor_get(x_710, 0); lean_inc(x_711); lean_dec_ref(x_710); -x_712 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_686, x_690, x_692, x_698, x_681, x_693, x_697, x_688, x_682, x_691, x_684); +x_712 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_692, x_693, x_694, x_697, x_690, x_686, x_695, x_681, x_696, x_688, x_691); if (lean_obj_tag(x_712) == 0) { lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; uint8_t x_717; @@ -27038,16 +27038,16 @@ lean_dec(x_721); x_723 = l_Lean_instToExprInt_mkNat(x_722); x_724 = l_Lean_mkApp3(x_718, x_719, x_720, x_723); x_651 = x_715; -x_652 = x_711; -x_653 = lean_box(0); -x_654 = x_694; -x_655 = x_701; +x_652 = x_713; +x_653 = x_683; +x_654 = x_685; +x_655 = lean_box(0); x_656 = x_714; -x_657 = x_713; -x_658 = x_696; -x_659 = x_709; +x_657 = x_701; +x_658 = x_709; +x_659 = x_705; x_660 = x_703; -x_661 = x_705; +x_661 = x_711; x_662 = x_724; goto block_680; } @@ -27058,16 +27058,16 @@ x_725 = l_Int_toNat(x_706); lean_dec(x_706); x_726 = l_Lean_instToExprInt_mkNat(x_725); x_651 = x_715; -x_652 = x_711; -x_653 = lean_box(0); -x_654 = x_694; -x_655 = x_701; +x_652 = x_713; +x_653 = x_683; +x_654 = x_685; +x_655 = lean_box(0); x_656 = x_714; -x_657 = x_713; -x_658 = x_696; -x_659 = x_709; +x_657 = x_701; +x_658 = x_709; +x_659 = x_705; x_660 = x_703; -x_661 = x_705; +x_661 = x_711; x_662 = x_726; goto block_680; } @@ -27081,8 +27081,8 @@ lean_dec(x_705); lean_dec(x_703); lean_dec(x_701); lean_dec(x_699); -lean_dec(x_696); -lean_dec(x_694); +lean_dec(x_685); +lean_dec(x_683); return x_712; } } @@ -27090,7 +27090,7 @@ return x_712; { lean_object* x_739; lean_object* x_740; lean_object* x_741; x_739 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_740 = l_Lean_mkApp10(x_734, x_731, x_728, x_729, x_737, x_736, x_735, x_738, x_730, x_732, x_739); +x_740 = l_Lean_mkApp10(x_731, x_729, x_732, x_736, x_735, x_728, x_734, x_738, x_730, x_733, x_739); x_741 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_741, 0, x_740); return x_741; @@ -27098,60 +27098,60 @@ return x_741; block_772: { lean_object* x_755; lean_object* x_756; uint8_t x_757; -x_755 = l_Lean_mkNatLit(x_749); +x_755 = l_Lean_mkNatLit(x_750); x_756 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_757 = lean_int_dec_le(x_756, x_748); +x_757 = lean_int_dec_le(x_756, x_746); if (x_757 == 0) { lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; x_758 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_759 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_746); +lean_inc(x_749); x_760 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_760, 0, x_759); -lean_ctor_set(x_760, 1, x_746); +lean_ctor_set(x_760, 1, x_749); x_761 = l_Lean_Expr_const___override(x_758, x_760); x_762 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__11; -lean_inc(x_746); -x_763 = l_Lean_Expr_const___override(x_762, x_746); +lean_inc(x_749); +x_763 = l_Lean_Expr_const___override(x_762, x_749); x_764 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__14; -x_765 = l_Lean_Expr_const___override(x_764, x_746); -x_766 = lean_int_neg(x_748); -lean_dec(x_748); +x_765 = l_Lean_Expr_const___override(x_764, x_749); +x_766 = lean_int_neg(x_746); +lean_dec(x_746); x_767 = l_Int_toNat(x_766); lean_dec(x_766); x_768 = l_Lean_instToExprInt_mkNat(x_767); x_769 = l_Lean_mkApp3(x_761, x_763, x_765, x_768); -x_728 = x_743; -x_729 = x_744; +x_728 = x_754; +x_729 = x_743; x_730 = x_745; -x_731 = x_747; -x_732 = x_750; -x_733 = lean_box(0); -x_734 = x_751; -x_735 = x_755; -x_736 = x_754; -x_737 = x_753; +x_731 = x_744; +x_732 = x_747; +x_733 = x_748; +x_734 = x_755; +x_735 = x_752; +x_736 = x_751; +x_737 = lean_box(0); x_738 = x_769; goto block_742; } else { lean_object* x_770; lean_object* x_771; +lean_dec(x_749); +x_770 = l_Int_toNat(x_746); lean_dec(x_746); -x_770 = l_Int_toNat(x_748); -lean_dec(x_748); x_771 = l_Lean_instToExprInt_mkNat(x_770); -x_728 = x_743; -x_729 = x_744; +x_728 = x_754; +x_729 = x_743; x_730 = x_745; -x_731 = x_747; -x_732 = x_750; -x_733 = lean_box(0); -x_734 = x_751; -x_735 = x_755; -x_736 = x_754; -x_737 = x_753; +x_731 = x_744; +x_732 = x_747; +x_733 = x_748; +x_734 = x_755; +x_735 = x_752; +x_736 = x_751; +x_737 = lean_box(0); x_738 = x_771; goto block_742; } @@ -27159,32 +27159,32 @@ goto block_742; block_819: { lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; -x_792 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_784); +x_792 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_getContext___redArg(x_782); x_793 = lean_ctor_get(x_792, 0); lean_inc(x_793); lean_dec_ref(x_792); -x_794 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_778, x_784, x_786, x_790, x_773, x_787, x_789, x_779, x_774, x_785, x_776); +x_794 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_784, x_782, x_783, x_789, x_780, x_777, x_786, x_773, x_787, x_779, x_781); x_795 = lean_ctor_get(x_794, 0); lean_inc(x_795); lean_dec_ref(x_794); -x_796 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_782, x_784, x_786, x_790, x_773, x_787, x_789, x_779, x_774, x_785, x_776); +x_796 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_774, x_782, x_783, x_789, x_780, x_777, x_786, x_773, x_787, x_779, x_781); x_797 = lean_ctor_get(x_796, 0); lean_inc(x_797); lean_dec_ref(x_796); -x_798 = lean_ctor_get(x_781, 0); +x_798 = lean_ctor_get(x_785, 0); lean_inc(x_798); -x_799 = lean_ctor_get(x_781, 1); +x_799 = lean_ctor_get(x_785, 1); lean_inc_ref(x_799); -lean_dec_ref(x_781); -x_800 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_799, x_784, x_786, x_790, x_773, x_787, x_789, x_779, x_774, x_785, x_776); +lean_dec_ref(x_785); +x_800 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_799, x_782, x_783, x_789, x_780, x_777, x_786, x_773, x_787, x_779, x_781); x_801 = lean_ctor_get(x_800, 0); lean_inc(x_801); lean_dec_ref(x_800); -x_802 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_777, x_784, x_786, x_790, x_773, x_787, x_789, x_779, x_774, x_785, x_776); +x_802 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkPolyDecl(x_775, x_782, x_783, x_789, x_780, x_777, x_786, x_773, x_787, x_779, x_781); x_803 = lean_ctor_get(x_802, 0); lean_inc(x_803); lean_dec_ref(x_802); -x_804 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_783, x_784, x_786, x_790, x_773, x_787, x_789, x_779, x_774, x_785, x_776); +x_804 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_CooperSplit_toExprProof(x_788, x_782, x_783, x_789, x_780, x_777, x_786, x_773, x_787, x_779, x_781); if (lean_obj_tag(x_804) == 0) { lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; uint8_t x_809; @@ -27207,17 +27207,17 @@ x_814 = l_Int_toNat(x_813); lean_dec(x_813); x_815 = l_Lean_instToExprInt_mkNat(x_814); x_816 = l_Lean_mkApp3(x_810, x_811, x_812, x_815); -x_743 = x_795; -x_744 = x_797; +x_743 = x_793; +x_744 = x_807; x_745 = x_803; -x_746 = x_806; -x_747 = x_793; -x_748 = x_788; -x_749 = x_780; -x_750 = x_805; -x_751 = x_807; -x_752 = lean_box(0); -x_753 = x_801; +x_746 = x_776; +x_747 = x_795; +x_748 = x_805; +x_749 = x_806; +x_750 = x_778; +x_751 = x_797; +x_752 = x_801; +x_753 = lean_box(0); x_754 = x_816; goto block_772; } @@ -27227,17 +27227,17 @@ lean_object* x_817; lean_object* x_818; x_817 = l_Int_toNat(x_798); lean_dec(x_798); x_818 = l_Lean_instToExprInt_mkNat(x_817); -x_743 = x_795; -x_744 = x_797; +x_743 = x_793; +x_744 = x_807; x_745 = x_803; -x_746 = x_806; -x_747 = x_793; -x_748 = x_788; -x_749 = x_780; -x_750 = x_805; -x_751 = x_807; -x_752 = lean_box(0); -x_753 = x_801; +x_746 = x_776; +x_747 = x_795; +x_748 = x_805; +x_749 = x_806; +x_750 = x_778; +x_751 = x_797; +x_752 = x_801; +x_753 = lean_box(0); x_754 = x_818; goto block_772; } @@ -27251,15 +27251,15 @@ lean_dec(x_797); lean_dec(x_795); lean_dec(x_793); lean_dec(x_791); -lean_dec(x_788); -lean_dec(x_780); +lean_dec(x_778); +lean_dec(x_776); return x_804; } } block_830: { lean_object* x_828; lean_object* x_829; -x_828 = l_Lean_mkApp6(x_824, x_823, x_827, x_822, x_820, x_821, x_826); +x_828 = l_Lean_mkApp6(x_823, x_826, x_827, x_822, x_821, x_820, x_825); x_829 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_829, 0, x_828); return x_829; @@ -27383,12 +27383,12 @@ x_871 = l_Int_toNat(x_870); lean_dec(x_870); x_872 = l_Lean_instToExprInt_mkNat(x_871); x_873 = l_Lean_mkApp3(x_867, x_868, x_869, x_872); -x_14 = x_861; -x_15 = lean_box(0); -x_16 = x_857; +x_14 = lean_box(0); +x_15 = x_862; +x_16 = x_859; x_17 = x_863; -x_18 = x_862; -x_19 = x_859; +x_18 = x_861; +x_19 = x_857; x_20 = x_873; goto block_24; } @@ -27398,12 +27398,12 @@ lean_object* x_874; lean_object* x_875; x_874 = l_Int_toNat(x_864); lean_dec(x_864); x_875 = l_Lean_instToExprInt_mkNat(x_874); -x_14 = x_861; -x_15 = lean_box(0); -x_16 = x_857; +x_14 = lean_box(0); +x_15 = x_862; +x_16 = x_859; x_17 = x_863; -x_18 = x_862; -x_19 = x_859; +x_18 = x_861; +x_19 = x_857; x_20 = x_875; goto block_24; } @@ -27494,12 +27494,12 @@ x_895 = l_Int_toNat(x_894); lean_dec(x_894); x_896 = l_Lean_instToExprInt_mkNat(x_895); x_897 = l_Lean_mkApp3(x_891, x_892, x_893, x_896); -x_25 = x_879; -x_26 = x_888; +x_25 = lean_box(0); +x_26 = x_885; x_27 = x_887; -x_28 = x_885; -x_29 = lean_box(0); -x_30 = x_883; +x_28 = x_883; +x_29 = x_879; +x_30 = x_888; x_31 = x_897; goto block_35; } @@ -27509,12 +27509,12 @@ lean_object* x_898; lean_object* x_899; x_898 = l_Int_toNat(x_880); lean_dec(x_880); x_899 = l_Lean_instToExprInt_mkNat(x_898); -x_25 = x_879; -x_26 = x_888; +x_25 = lean_box(0); +x_26 = x_885; x_27 = x_887; -x_28 = x_885; -x_29 = lean_box(0); -x_30 = x_883; +x_28 = x_883; +x_29 = x_879; +x_30 = x_888; x_31 = x_899; goto block_35; } @@ -27548,22 +27548,22 @@ x_906 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; x_907 = lean_int_dec_lt(x_903, x_906); if (x_907 == 0) { -x_107 = x_834; -x_108 = x_903; -x_109 = x_838; -x_110 = lean_box(0); -x_111 = x_840; -x_112 = x_837; -x_113 = x_902; -x_114 = x_904; +x_107 = x_837; +x_108 = x_902; +x_109 = x_904; +x_110 = x_901; +x_111 = x_835; +x_112 = x_839; +x_113 = x_834; +x_114 = x_840; x_115 = x_831; -x_116 = x_839; -x_117 = x_900; -x_118 = x_832; -x_119 = x_835; -x_120 = x_901; -x_121 = x_836; -x_122 = x_833; +x_116 = x_900; +x_117 = x_832; +x_118 = x_903; +x_119 = x_836; +x_120 = x_838; +x_121 = x_833; +x_122 = lean_box(0); x_123 = x_905; goto block_146; } @@ -27572,22 +27572,22 @@ else lean_object* x_908; x_908 = lean_int_neg(x_905); lean_dec(x_905); -x_107 = x_834; -x_108 = x_903; -x_109 = x_838; -x_110 = lean_box(0); -x_111 = x_840; -x_112 = x_837; -x_113 = x_902; -x_114 = x_904; +x_107 = x_837; +x_108 = x_902; +x_109 = x_904; +x_110 = x_901; +x_111 = x_835; +x_112 = x_839; +x_113 = x_834; +x_114 = x_840; x_115 = x_831; -x_116 = x_839; -x_117 = x_900; -x_118 = x_832; -x_119 = x_835; -x_120 = x_901; -x_121 = x_836; -x_122 = x_833; +x_116 = x_900; +x_117 = x_832; +x_118 = x_903; +x_119 = x_836; +x_120 = x_838; +x_121 = x_833; +x_122 = lean_box(0); x_123 = x_908; goto block_146; } @@ -27690,31 +27690,31 @@ x_949 = l_Int_toNat(x_948); lean_dec(x_948); x_950 = l_Lean_instToExprInt_mkNat(x_949); x_951 = l_Lean_mkApp3(x_945, x_946, x_947, x_950); -x_361 = x_834; -x_362 = x_934; -x_363 = x_840; -x_364 = x_837; -x_365 = x_929; -x_366 = x_940; +x_361 = x_909; +x_362 = x_912; +x_363 = x_941; +x_364 = x_831; +x_365 = x_836; +x_366 = x_925; x_367 = x_937; -x_368 = x_941; +x_368 = x_837; x_369 = x_911; -x_370 = x_839; -x_371 = x_835; -x_372 = x_925; -x_373 = x_912; -x_374 = x_909; -x_375 = x_833; -x_376 = x_942; -x_377 = x_926; -x_378 = x_838; -x_379 = lean_box(0); -x_380 = x_933; -x_381 = x_939; -x_382 = x_831; -x_383 = x_927; -x_384 = x_832; -x_385 = x_836; +x_370 = x_835; +x_371 = x_934; +x_372 = x_926; +x_373 = x_942; +x_374 = x_940; +x_375 = x_927; +x_376 = x_933; +x_377 = x_839; +x_378 = lean_box(0); +x_379 = x_840; +x_380 = x_834; +x_381 = x_929; +x_382 = x_832; +x_383 = x_939; +x_384 = x_838; +x_385 = x_833; x_386 = x_951; goto block_404; } @@ -27723,31 +27723,31 @@ else lean_object* x_952; lean_object* x_953; x_952 = l_Int_toNat(x_930); x_953 = l_Lean_instToExprInt_mkNat(x_952); -x_361 = x_834; -x_362 = x_934; -x_363 = x_840; -x_364 = x_837; -x_365 = x_929; -x_366 = x_940; +x_361 = x_909; +x_362 = x_912; +x_363 = x_941; +x_364 = x_831; +x_365 = x_836; +x_366 = x_925; x_367 = x_937; -x_368 = x_941; +x_368 = x_837; x_369 = x_911; -x_370 = x_839; -x_371 = x_835; -x_372 = x_925; -x_373 = x_912; -x_374 = x_909; -x_375 = x_833; -x_376 = x_942; -x_377 = x_926; -x_378 = x_838; -x_379 = lean_box(0); -x_380 = x_933; -x_381 = x_939; -x_382 = x_831; -x_383 = x_927; -x_384 = x_832; -x_385 = x_836; +x_370 = x_835; +x_371 = x_934; +x_372 = x_926; +x_373 = x_942; +x_374 = x_940; +x_375 = x_927; +x_376 = x_933; +x_377 = x_839; +x_378 = lean_box(0); +x_379 = x_840; +x_380 = x_834; +x_381 = x_929; +x_382 = x_832; +x_383 = x_939; +x_384 = x_838; +x_385 = x_833; x_386 = x_953; goto block_404; } @@ -27901,18 +27901,18 @@ x_989 = l_Int_toNat(x_988); lean_dec(x_988); x_990 = l_Lean_instToExprInt_mkNat(x_989); x_991 = l_Lean_mkApp3(x_985, x_986, x_987, x_990); -x_451 = x_970; -x_452 = x_975; -x_453 = x_965; -x_454 = x_981; -x_455 = x_982; -x_456 = x_979; -x_457 = x_969; -x_458 = x_973; -x_459 = x_960; -x_460 = lean_box(0); -x_461 = x_977; -x_462 = x_980; +x_451 = x_975; +x_452 = x_973; +x_453 = x_960; +x_454 = x_970; +x_455 = x_965; +x_456 = x_969; +x_457 = x_981; +x_458 = x_979; +x_459 = x_977; +x_460 = x_982; +x_461 = x_980; +x_462 = lean_box(0); x_463 = x_991; goto block_481; } @@ -27922,18 +27922,18 @@ lean_object* x_992; lean_object* x_993; x_992 = l_Int_toNat(x_966); lean_dec(x_966); x_993 = l_Lean_instToExprInt_mkNat(x_992); -x_451 = x_970; -x_452 = x_975; -x_453 = x_965; -x_454 = x_981; -x_455 = x_982; -x_456 = x_979; -x_457 = x_969; -x_458 = x_973; -x_459 = x_960; -x_460 = lean_box(0); -x_461 = x_977; -x_462 = x_980; +x_451 = x_975; +x_452 = x_973; +x_453 = x_960; +x_454 = x_970; +x_455 = x_965; +x_456 = x_969; +x_457 = x_981; +x_458 = x_979; +x_459 = x_977; +x_460 = x_982; +x_461 = x_980; +x_462 = lean_box(0); x_463 = x_993; goto block_481; } @@ -28026,15 +28026,15 @@ x_1016 = l_Int_toNat(x_1015); lean_dec(x_1015); x_1017 = l_Lean_instToExprInt_mkNat(x_1016); x_1018 = l_Lean_mkApp3(x_1012, x_1013, x_1014, x_1017); -x_494 = x_1008; +x_494 = x_1007; x_495 = x_998; -x_496 = lean_box(0); -x_497 = x_1009; -x_498 = x_1002; -x_499 = x_1006; -x_500 = x_994; -x_501 = x_1007; -x_502 = x_1004; +x_496 = x_1008; +x_497 = x_1002; +x_498 = x_1004; +x_499 = x_994; +x_500 = x_1006; +x_501 = lean_box(0); +x_502 = x_1009; x_503 = x_1018; goto block_521; } @@ -28044,15 +28044,15 @@ lean_object* x_1019; lean_object* x_1020; x_1019 = l_Int_toNat(x_999); lean_dec(x_999); x_1020 = l_Lean_instToExprInt_mkNat(x_1019); -x_494 = x_1008; +x_494 = x_1007; x_495 = x_998; -x_496 = lean_box(0); -x_497 = x_1009; -x_498 = x_1002; -x_499 = x_1006; -x_500 = x_994; -x_501 = x_1007; -x_502 = x_1004; +x_496 = x_1008; +x_497 = x_1002; +x_498 = x_1004; +x_499 = x_994; +x_500 = x_1006; +x_501 = lean_box(0); +x_502 = x_1009; x_503 = x_1020; goto block_521; } @@ -28121,13 +28121,13 @@ x_1043 = l_Int_toNat(x_1042); lean_dec(x_1042); x_1044 = l_Lean_instToExprInt_mkNat(x_1043); x_1045 = l_Lean_mkApp3(x_1039, x_1040, x_1041, x_1044); -x_522 = x_1031; +x_522 = x_1026; x_523 = x_1028; -x_524 = x_1026; -x_525 = lean_box(0); -x_526 = x_1035; +x_524 = x_1036; +x_525 = x_1035; +x_526 = x_1031; x_527 = x_1033; -x_528 = x_1036; +x_528 = lean_box(0); x_529 = x_1045; goto block_533; } @@ -28137,13 +28137,13 @@ lean_object* x_1046; lean_object* x_1047; x_1046 = l_Int_toNat(x_1021); lean_dec(x_1021); x_1047 = l_Lean_instToExprInt_mkNat(x_1046); -x_522 = x_1031; +x_522 = x_1026; x_523 = x_1028; -x_524 = x_1026; -x_525 = lean_box(0); -x_526 = x_1035; +x_524 = x_1036; +x_525 = x_1035; +x_526 = x_1031; x_527 = x_1033; -x_528 = x_1036; +x_528 = lean_box(0); x_529 = x_1047; goto block_533; } @@ -28242,17 +28242,17 @@ lean_dec(x_1078); x_1080 = l_Lean_instToExprInt_mkNat(x_1079); x_1081 = l_Lean_mkApp3(x_1075, x_1076, x_1077, x_1080); x_549 = x_1059; -x_550 = x_1067; -x_551 = x_1065; -x_552 = x_1071; -x_553 = x_1056; -x_554 = x_1048; -x_555 = x_1054; +x_550 = lean_box(0); +x_551 = x_1071; +x_552 = x_1065; +x_553 = x_1048; +x_554 = x_1072; +x_555 = x_1070; x_556 = x_1063; -x_557 = lean_box(0); -x_558 = x_1070; -x_559 = x_1072; -x_560 = x_1069; +x_557 = x_1069; +x_558 = x_1054; +x_559 = x_1056; +x_560 = x_1067; x_561 = x_1081; goto block_579; } @@ -28263,17 +28263,17 @@ x_1082 = l_Int_toNat(x_1060); lean_dec(x_1060); x_1083 = l_Lean_instToExprInt_mkNat(x_1082); x_549 = x_1059; -x_550 = x_1067; -x_551 = x_1065; -x_552 = x_1071; -x_553 = x_1056; -x_554 = x_1048; -x_555 = x_1054; +x_550 = lean_box(0); +x_551 = x_1071; +x_552 = x_1065; +x_553 = x_1048; +x_554 = x_1072; +x_555 = x_1070; x_556 = x_1063; -x_557 = lean_box(0); -x_558 = x_1070; -x_559 = x_1072; -x_560 = x_1069; +x_557 = x_1069; +x_558 = x_1054; +x_559 = x_1056; +x_560 = x_1067; x_561 = x_1083; goto block_579; } @@ -28342,23 +28342,23 @@ lean_inc_ref(x_1093); x_1094 = lean_ctor_get(x_1089, 0); lean_inc_ref(x_1094); x_1095 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__25; -x_593 = x_834; -x_594 = x_838; -x_595 = lean_box(0); -x_596 = x_840; -x_597 = x_1094; -x_598 = x_1084; -x_599 = x_1091; -x_600 = x_837; -x_601 = x_831; -x_602 = x_839; -x_603 = x_832; -x_604 = x_835; -x_605 = x_1092; -x_606 = x_1093; -x_607 = x_1090; -x_608 = x_836; -x_609 = x_833; +x_593 = x_837; +x_594 = x_1094; +x_595 = x_1092; +x_596 = x_1091; +x_597 = x_1090; +x_598 = x_835; +x_599 = x_839; +x_600 = x_1093; +x_601 = x_834; +x_602 = x_840; +x_603 = x_1084; +x_604 = x_831; +x_605 = x_832; +x_606 = x_836; +x_607 = x_838; +x_608 = x_833; +x_609 = lean_box(0); x_610 = x_1095; goto block_635; } @@ -28379,23 +28379,23 @@ lean_inc_ref(x_1101); x_1102 = lean_ctor_get(x_1097, 0); lean_inc_ref(x_1102); x_1103 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__27; -x_593 = x_834; -x_594 = x_838; -x_595 = lean_box(0); -x_596 = x_840; -x_597 = x_1102; -x_598 = x_1084; -x_599 = x_1099; -x_600 = x_837; -x_601 = x_831; -x_602 = x_839; -x_603 = x_832; -x_604 = x_835; -x_605 = x_1100; -x_606 = x_1101; -x_607 = x_1098; -x_608 = x_836; -x_609 = x_833; +x_593 = x_837; +x_594 = x_1102; +x_595 = x_1100; +x_596 = x_1099; +x_597 = x_1098; +x_598 = x_835; +x_599 = x_839; +x_600 = x_1101; +x_601 = x_834; +x_602 = x_840; +x_603 = x_1084; +x_604 = x_831; +x_605 = x_832; +x_606 = x_836; +x_607 = x_838; +x_608 = x_833; +x_609 = lean_box(0); x_610 = x_1103; goto block_635; } @@ -28423,24 +28423,24 @@ lean_inc_ref(x_1111); x_1112 = lean_ctor_get(x_1086, 0); lean_inc(x_1112); x_1113 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__29; -x_681 = x_834; -x_682 = x_838; -x_683 = lean_box(0); -x_684 = x_840; -x_685 = x_1111; -x_686 = x_1084; -x_687 = x_1108; -x_688 = x_837; -x_689 = x_1112; -x_690 = x_831; -x_691 = x_839; -x_692 = x_832; -x_693 = x_835; -x_694 = x_1109; -x_695 = x_1110; -x_696 = x_1107; -x_697 = x_836; -x_698 = x_833; +x_681 = x_837; +x_682 = x_1111; +x_683 = x_1109; +x_684 = x_1108; +x_685 = x_1107; +x_686 = x_835; +x_687 = x_1112; +x_688 = x_839; +x_689 = x_1110; +x_690 = x_834; +x_691 = x_840; +x_692 = x_1084; +x_693 = x_831; +x_694 = x_832; +x_695 = x_836; +x_696 = x_838; +x_697 = x_833; +x_698 = lean_box(0); x_699 = x_1113; goto block_727; } @@ -28463,24 +28463,24 @@ lean_inc_ref(x_1120); x_1121 = lean_ctor_get(x_1086, 0); lean_inc(x_1121); x_1122 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__31; -x_681 = x_834; -x_682 = x_838; -x_683 = lean_box(0); -x_684 = x_840; -x_685 = x_1120; -x_686 = x_1084; -x_687 = x_1117; -x_688 = x_837; -x_689 = x_1121; -x_690 = x_831; -x_691 = x_839; -x_692 = x_832; -x_693 = x_835; -x_694 = x_1118; -x_695 = x_1119; -x_696 = x_1116; -x_697 = x_836; -x_698 = x_833; +x_681 = x_837; +x_682 = x_1120; +x_683 = x_1118; +x_684 = x_1117; +x_685 = x_1116; +x_686 = x_835; +x_687 = x_1121; +x_688 = x_839; +x_689 = x_1119; +x_690 = x_834; +x_691 = x_840; +x_692 = x_1084; +x_693 = x_831; +x_694 = x_832; +x_695 = x_836; +x_696 = x_838; +x_697 = x_833; +x_698 = lean_box(0); x_699 = x_1122; goto block_727; } @@ -28535,24 +28535,24 @@ lean_inc_ref(x_1135); x_1136 = lean_ctor_get(x_1130, 0); lean_inc_ref(x_1136); x_1137 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__35; -x_773 = x_834; -x_774 = x_838; -x_775 = lean_box(0); -x_776 = x_840; -x_777 = x_1132; -x_778 = x_1135; -x_779 = x_837; -x_780 = x_1133; -x_781 = x_1134; -x_782 = x_1136; -x_783 = x_1123; -x_784 = x_831; -x_785 = x_839; -x_786 = x_832; -x_787 = x_835; -x_788 = x_1131; -x_789 = x_836; -x_790 = x_833; +x_773 = x_837; +x_774 = x_1136; +x_775 = x_1132; +x_776 = x_1131; +x_777 = x_835; +x_778 = x_1133; +x_779 = x_839; +x_780 = x_834; +x_781 = x_840; +x_782 = x_831; +x_783 = x_832; +x_784 = x_1135; +x_785 = x_1134; +x_786 = x_836; +x_787 = x_838; +x_788 = x_1123; +x_789 = x_833; +x_790 = lean_box(0); x_791 = x_1137; goto block_819; } @@ -28575,24 +28575,24 @@ lean_inc_ref(x_1144); x_1145 = lean_ctor_get(x_1139, 0); lean_inc_ref(x_1145); x_1146 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_DvdCnstr_toExprProof___lam__0___closed__37; -x_773 = x_834; -x_774 = x_838; -x_775 = lean_box(0); -x_776 = x_840; -x_777 = x_1141; -x_778 = x_1144; -x_779 = x_837; -x_780 = x_1142; -x_781 = x_1143; -x_782 = x_1145; -x_783 = x_1123; -x_784 = x_831; -x_785 = x_839; -x_786 = x_832; -x_787 = x_835; -x_788 = x_1140; -x_789 = x_836; -x_790 = x_833; +x_773 = x_837; +x_774 = x_1145; +x_775 = x_1141; +x_776 = x_1140; +x_777 = x_835; +x_778 = x_1142; +x_779 = x_839; +x_780 = x_834; +x_781 = x_840; +x_782 = x_831; +x_783 = x_832; +x_784 = x_1144; +x_785 = x_1143; +x_786 = x_836; +x_787 = x_838; +x_788 = x_1123; +x_789 = x_833; +x_790 = lean_box(0); x_791 = x_1146; goto block_819; } @@ -28678,13 +28678,13 @@ x_1180 = l_Int_toNat(x_1179); lean_dec(x_1179); x_1181 = l_Lean_instToExprInt_mkNat(x_1180); x_1182 = l_Lean_mkApp3(x_1176, x_1177, x_1178, x_1181); -x_820 = x_1167; -x_821 = x_1172; +x_820 = x_1172; +x_821 = x_1167; x_822 = x_1165; -x_823 = x_1161; -x_824 = x_1173; -x_825 = lean_box(0); -x_826 = x_1169; +x_823 = x_1173; +x_824 = lean_box(0); +x_825 = x_1169; +x_826 = x_1161; x_827 = x_1182; goto block_830; } @@ -28694,13 +28694,13 @@ lean_object* x_1183; lean_object* x_1184; x_1183 = l_Int_toNat(x_1162); lean_dec(x_1162); x_1184 = l_Lean_instToExprInt_mkNat(x_1183); -x_820 = x_1167; -x_821 = x_1172; +x_820 = x_1172; +x_821 = x_1167; x_822 = x_1165; -x_823 = x_1161; -x_824 = x_1173; -x_825 = lean_box(0); -x_826 = x_1169; +x_823 = x_1173; +x_824 = lean_box(0); +x_825 = x_1169; +x_826 = x_1161; x_827 = x_1184; goto block_830; } @@ -30071,7 +30071,7 @@ return x_429; { lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_22 = l_Lean_mkApp6(x_19, x_17, x_15, x_16, x_20, x_21, x_14); +x_22 = l_Lean_mkApp6(x_18, x_14, x_19, x_16, x_20, x_21, x_17); x_23 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_23, 0, x_22); return x_23; @@ -30080,7 +30080,7 @@ return x_23; { lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__4; -x_36 = l_Lean_mkApp9(x_29, x_33, x_30, x_34, x_32, x_31, x_28, x_35, x_26, x_25); +x_36 = l_Lean_mkApp9(x_26, x_25, x_30, x_34, x_33, x_32, x_27, x_35, x_31, x_29); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); return x_37; @@ -30089,38 +30089,38 @@ return x_37; { lean_object* x_51; uint8_t x_52; x_51 = l_Lean_Meta_Grind_Arith_Cutsat_instHashableExpr__lean_hash___closed__0; -x_52 = lean_int_dec_le(x_51, x_48); +x_52 = lean_int_dec_le(x_51, x_44); if (x_52 == 0) { lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_53 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__7; x_54 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__8; -lean_inc(x_47); +lean_inc(x_43); x_55 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_47); +lean_ctor_set(x_55, 1, x_43); x_56 = l_Lean_Expr_const___override(x_53, x_55); -lean_inc_ref(x_40); -x_57 = l_Lean_Name_mkStr1(x_40); -lean_inc(x_47); -x_58 = l_Lean_Expr_const___override(x_57, x_47); +lean_inc_ref(x_41); +x_57 = l_Lean_Name_mkStr1(x_41); +lean_inc(x_43); +x_58 = l_Lean_Expr_const___override(x_57, x_43); x_59 = l___private_Lean_Meta_Tactic_Grind_Arith_Cutsat_Proof_0__Lean_Meta_Grind_Arith_Cutsat_mkMulEqProof_goVar___closed__13; -x_60 = l_Lean_Name_mkStr2(x_40, x_59); -x_61 = l_Lean_Expr_const___override(x_60, x_47); -x_62 = lean_int_neg(x_48); -lean_dec(x_48); +x_60 = l_Lean_Name_mkStr2(x_41, x_59); +x_61 = l_Lean_Expr_const___override(x_60, x_43); +x_62 = lean_int_neg(x_44); +lean_dec(x_44); x_63 = l_Int_toNat(x_62); lean_dec(x_62); x_64 = l_Lean_instToExprInt_mkNat(x_63); x_65 = l_Lean_mkApp3(x_56, x_58, x_61, x_64); x_25 = x_39; -x_26 = x_41; -x_27 = lean_box(0); -x_28 = x_44; -x_29 = x_43; +x_26 = x_40; +x_27 = x_42; +x_28 = lean_box(0); +x_29 = x_45; x_30 = x_50; -x_31 = x_45; -x_32 = x_46; +x_31 = x_48; +x_32 = x_47; x_33 = x_49; x_34 = x_65; goto block_38; @@ -30128,19 +30128,19 @@ goto block_38; else { lean_object* x_66; lean_object* x_67; -lean_dec(x_47); -lean_dec_ref(x_40); -x_66 = l_Int_toNat(x_48); -lean_dec(x_48); +lean_dec(x_43); +lean_dec_ref(x_41); +x_66 = l_Int_toNat(x_44); +lean_dec(x_44); x_67 = l_Lean_instToExprInt_mkNat(x_66); x_25 = x_39; -x_26 = x_41; -x_27 = lean_box(0); -x_28 = x_44; -x_29 = x_43; +x_26 = x_40; +x_27 = x_42; +x_28 = lean_box(0); +x_29 = x_45; x_30 = x_50; -x_31 = x_45; -x_32 = x_46; +x_31 = x_48; +x_32 = x_47; x_33 = x_49; x_34 = x_67; goto block_38; @@ -30651,12 +30651,12 @@ x_223 = l_Int_toNat(x_222); lean_dec(x_222); x_224 = l_Lean_instToExprInt_mkNat(x_223); x_225 = l_Lean_mkApp3(x_219, x_220, x_221, x_224); -x_14 = x_213; -x_15 = x_209; +x_14 = x_206; +x_15 = lean_box(0); x_16 = x_211; -x_17 = x_206; -x_18 = lean_box(0); -x_19 = x_216; +x_17 = x_213; +x_18 = x_216; +x_19 = x_209; x_20 = x_225; goto block_24; } @@ -30666,12 +30666,12 @@ lean_object* x_226; lean_object* x_227; x_226 = l_Int_toNat(x_215); lean_dec(x_215); x_227 = l_Lean_instToExprInt_mkNat(x_226); -x_14 = x_213; -x_15 = x_209; +x_14 = x_206; +x_15 = lean_box(0); x_16 = x_211; -x_17 = x_206; -x_18 = lean_box(0); -x_19 = x_216; +x_17 = x_213; +x_18 = x_216; +x_19 = x_209; x_20 = x_227; goto block_24; } @@ -30767,17 +30767,17 @@ x_257 = l_Int_toNat(x_256); lean_dec(x_256); x_258 = l_Lean_instToExprInt_mkNat(x_257); x_259 = l_Lean_mkApp3(x_253, x_254, x_255, x_258); -x_39 = x_245; -x_40 = x_248; -x_41 = x_243; -x_42 = lean_box(0); -x_43 = x_250; -x_44 = x_241; -x_45 = x_239; -x_46 = x_236; -x_47 = x_249; -x_48 = x_247; -x_49 = x_233; +x_39 = x_233; +x_40 = x_250; +x_41 = x_248; +x_42 = x_241; +x_43 = x_249; +x_44 = x_247; +x_45 = x_245; +x_46 = lean_box(0); +x_47 = x_239; +x_48 = x_243; +x_49 = x_236; x_50 = x_259; goto block_68; } @@ -30787,17 +30787,17 @@ lean_object* x_260; lean_object* x_261; x_260 = l_Int_toNat(x_246); lean_dec(x_246); x_261 = l_Lean_instToExprInt_mkNat(x_260); -x_39 = x_245; -x_40 = x_248; -x_41 = x_243; -x_42 = lean_box(0); -x_43 = x_250; -x_44 = x_241; -x_45 = x_239; -x_46 = x_236; -x_47 = x_249; -x_48 = x_247; -x_49 = x_233; +x_39 = x_233; +x_40 = x_250; +x_41 = x_248; +x_42 = x_241; +x_43 = x_249; +x_44 = x_247; +x_45 = x_245; +x_46 = lean_box(0); +x_47 = x_239; +x_48 = x_243; +x_49 = x_236; x_50 = x_261; goto block_68; } diff --git a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c index c8176d77b4cf..88baf4c77911 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c +++ b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c @@ -10562,8 +10562,8 @@ goto block_106; { lean_object* x_14; lean_object* x_15; x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_11); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_12); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); return x_15; @@ -10577,9 +10577,9 @@ if (x_24 == 0) { if (x_1 == 0) { -lean_dec_ref(x_19); -x_11 = x_17; -x_12 = x_23; +lean_dec_ref(x_21); +x_11 = x_23; +x_12 = x_20; x_13 = lean_box(0); goto block_16; } @@ -10589,10 +10589,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_18, x_21, x_19, x_22); -lean_dec_ref(x_19); +x_26 = l_Lean_throwError___at___00Lean_Meta_SolveByElim_SolveByElimConfig_testPartialSolutions_spec__3___redArg(x_25, x_22, x_19, x_21, x_17); +lean_dec_ref(x_21); x_27 = !lean_is_exclusive(x_26); if (x_27 == 0) { @@ -10611,9 +10611,9 @@ return x_29; } else { -lean_dec_ref(x_19); -x_11 = x_17; -x_12 = x_23; +lean_dec_ref(x_21); +x_11 = x_23; +x_12 = x_20; x_13 = lean_box(0); goto block_16; } @@ -10621,9 +10621,9 @@ goto block_16; } else { -lean_dec_ref(x_19); -x_11 = x_17; -x_12 = x_23; +lean_dec_ref(x_21); +x_11 = x_23; +x_12 = x_20; x_13 = lean_box(0); goto block_16; } @@ -10632,35 +10632,35 @@ goto block_16; { lean_object* x_42; lean_object* x_43; x_42 = lean_array_to_list(x_41); -lean_inc(x_38); -x_43 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__5(x_42, x_38); +lean_inc(x_35); +x_43 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__5(x_42, x_35); 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__4(x_3, x_38); +x_44 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__4(x_3, x_35); x_45 = l_List_appendTR___redArg(x_44, x_43); -x_46 = l_List_appendTR___redArg(x_45, x_39); -x_17 = x_33; -x_18 = x_34; -x_19 = x_36; -x_20 = lean_box(0); -x_21 = x_37; -x_22 = x_40; +x_46 = l_List_appendTR___redArg(x_45, x_34); +x_17 = x_36; +x_18 = lean_box(0); +x_19 = x_38; +x_20 = x_33; +x_21 = x_40; +x_22 = x_39; x_23 = x_46; goto block_30; } else { lean_object* x_47; lean_object* x_48; -lean_dec(x_39); -x_47 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__4(x_3, x_38); +lean_dec(x_34); +x_47 = l_List_mapTR_loop___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__4(x_3, x_35); x_48 = l_List_appendTR___redArg(x_47, x_43); -x_17 = x_33; -x_18 = x_34; -x_19 = x_36; -x_20 = lean_box(0); -x_21 = x_37; -x_22 = x_40; +x_17 = x_36; +x_18 = lean_box(0); +x_19 = x_38; +x_20 = x_33; +x_21 = x_40; +x_22 = x_39; x_23 = x_48; goto block_30; } @@ -10760,13 +10760,13 @@ if (x_99 == 0) { lean_dec(x_98); lean_dec(x_64); -x_34 = x_50; -x_35 = lean_box(0); -x_36 = x_52; -x_37 = x_51; -x_38 = x_73; -x_39 = x_95; -x_40 = x_53; +x_34 = x_95; +x_35 = x_73; +x_36 = x_53; +x_37 = lean_box(0); +x_38 = x_51; +x_39 = x_50; +x_40 = x_52; x_41 = x_96; goto block_49; } @@ -10778,13 +10778,13 @@ if (x_100 == 0) { lean_dec(x_98); lean_dec(x_64); -x_34 = x_50; -x_35 = lean_box(0); -x_36 = x_52; -x_37 = x_51; -x_38 = x_73; -x_39 = x_95; -x_40 = x_53; +x_34 = x_95; +x_35 = x_73; +x_36 = x_53; +x_37 = lean_box(0); +x_38 = x_51; +x_39 = x_50; +x_40 = x_52; x_41 = x_96; goto block_49; } @@ -10795,13 +10795,13 @@ x_101 = lean_usize_of_nat(x_98); lean_dec(x_98); x_102 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Lean_Meta_SolveByElim_mkAssumptionSet_spec__6(x_64, x_62, x_101, x_96); lean_dec(x_64); -x_34 = x_50; -x_35 = lean_box(0); -x_36 = x_52; -x_37 = x_51; -x_38 = x_73; -x_39 = x_95; -x_40 = x_53; +x_34 = x_95; +x_35 = x_73; +x_36 = x_53; +x_37 = lean_box(0); +x_38 = x_51; +x_39 = x_50; +x_40 = x_52; x_41 = x_102; goto block_49; } diff --git a/stage0/stdlib/Lean/Widget/Commands.c b/stage0/stdlib/Lean/Widget/Commands.c index fef5752adc15..5bd7346a34f4 100644 --- a/stage0/stdlib/Lean/Widget/Commands.c +++ b/stage0/stdlib/Lean/Widget/Commands.c @@ -7975,9 +7975,9 @@ goto block_451; block_444: { lean_object* x_441; lean_object* x_442; lean_object* x_443; -x_441 = lean_nat_add(x_439, x_440); +x_441 = lean_nat_add(x_438, x_440); lean_dec(x_440); -lean_dec(x_439); +lean_dec(x_438); if (lean_is_scalar(x_435)) { x_442 = lean_alloc_ctor(0, 5, 0); } else { @@ -7996,7 +7996,7 @@ if (lean_is_scalar(x_425)) { lean_ctor_set(x_443, 0, x_437); lean_ctor_set(x_443, 1, x_427); lean_ctor_set(x_443, 2, x_428); -lean_ctor_set(x_443, 3, x_438); +lean_ctor_set(x_443, 3, x_439); lean_ctor_set(x_443, 4, x_442); return x_443; } @@ -8022,8 +8022,8 @@ if (lean_obj_tag(x_430) == 0) lean_object* x_449; x_449 = lean_ctor_get(x_430, 0); lean_inc(x_449); -x_438 = x_447; -x_439 = x_448; +x_438 = x_448; +x_439 = x_447; x_440 = x_449; goto block_444; } @@ -8031,8 +8031,8 @@ else { lean_object* x_450; x_450 = lean_unsigned_to_nat(0u); -x_438 = x_447; -x_439 = x_448; +x_438 = x_448; +x_439 = x_447; x_440 = x_450; goto block_444; } diff --git a/stage0/stdlib/Std/Data/DHashMap/Internal/RawLemmas.c b/stage0/stdlib/Std/Data/DHashMap/Internal/RawLemmas.c index ece4a1bbc431..3089a5d506e6 100644 --- a/stage0/stdlib/Std/Data/DHashMap/Internal/RawLemmas.c +++ b/stage0/stdlib/Std/Data/DHashMap/Internal/RawLemmas.c @@ -11354,7 +11354,7 @@ return x_2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; x_45 = l_Std_DHashMap_Internal_Raw_u2080_tacticWf__trivial___closed__0; x_46 = l_Std_DHashMap_Internal_Raw_u2080_tacticWf__trivial___closed__1; x_47 = l_Std_DHashMap_Internal_Raw_u2080_tacticWf__trivial___closed__2; @@ -11425,7 +11425,7 @@ goto block_245; block_133: { lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; size_t x_113; size_t x_114; lean_object* x_115; size_t x_116; lean_object* x_117; lean_object* x_118; size_t x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_57); +lean_dec(x_55); x_61 = lean_ctor_get(x_59, 5); lean_inc(x_61); lean_dec_ref(x_59); @@ -11531,9 +11531,9 @@ lean_inc(x_63); x_112 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_112, 0, x_63); lean_ctor_set(x_112, 1, x_111); -x_113 = lean_array_size(x_54); +x_113 = lean_array_size(x_57); x_114 = 0; -x_115 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__2(x_113, x_114, x_54); +x_115 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__2(x_113, x_114, x_57); x_116 = lean_array_size(x_115); x_117 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__3___redArg(x_116, x_114, x_115); x_118 = l_Array_append___redArg(x_117, x_58); @@ -11566,51 +11566,51 @@ x_129 = l_Array_mkArray2___redArg(x_128, x_79); if (lean_obj_tag(x_56) == 0) { x_4 = x_99; -x_5 = x_71; -x_6 = x_90; -x_7 = x_70; +x_5 = x_64; +x_6 = x_65; +x_7 = x_92; x_8 = x_60; -x_9 = x_92; -x_10 = x_66; -x_11 = x_72; -x_12 = x_64; -x_13 = x_65; -x_14 = x_67; -x_15 = x_63; -x_16 = x_105; -x_17 = x_68; -x_18 = x_73; -x_19 = x_129; -x_20 = x_114; -x_21 = x_55; +x_9 = x_73; +x_10 = x_71; +x_11 = x_114; +x_12 = x_72; +x_13 = x_90; +x_14 = x_63; +x_15 = x_67; +x_16 = x_70; +x_17 = x_66; +x_18 = x_129; +x_19 = x_105; +x_20 = x_68; +x_21 = x_54; goto block_44; } else { lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_dec_ref(x_55); +lean_dec_ref(x_54); x_130 = lean_ctor_get(x_56, 0); lean_inc(x_130); lean_dec_ref(x_56); x_131 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1___closed__17; x_132 = lean_array_push(x_131, x_130); x_4 = x_99; -x_5 = x_71; -x_6 = x_90; -x_7 = x_70; +x_5 = x_64; +x_6 = x_65; +x_7 = x_92; x_8 = x_60; -x_9 = x_92; -x_10 = x_66; -x_11 = x_72; -x_12 = x_64; -x_13 = x_65; -x_14 = x_67; -x_15 = x_63; -x_16 = x_105; -x_17 = x_68; -x_18 = x_73; -x_19 = x_129; -x_20 = x_114; +x_9 = x_73; +x_10 = x_71; +x_11 = x_114; +x_12 = x_72; +x_13 = x_90; +x_14 = x_63; +x_15 = x_67; +x_16 = x_70; +x_17 = x_66; +x_18 = x_129; +x_19 = x_105; +x_20 = x_68; x_21 = x_132; goto block_44; } @@ -11620,36 +11620,36 @@ goto block_44; size_t x_142; size_t x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; x_142 = lean_array_size(x_141); x_143 = 0; -lean_inc_ref(x_138); +lean_inc_ref(x_135); lean_inc_ref(x_136); -x_144 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__7(x_140, x_45, x_46, x_47, x_48, x_135, x_141, x_142, x_143, x_136, x_138, x_139); +x_144 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__7(x_137, x_45, x_46, x_47, x_48, x_140, x_141, x_142, x_143, x_136, x_135, x_134); lean_dec_ref(x_141); -lean_dec_ref(x_135); +lean_dec_ref(x_140); x_145 = lean_ctor_get(x_144, 0); lean_inc(x_145); x_146 = lean_ctor_get(x_144, 1); lean_inc(x_146); lean_dec_ref(x_144); -x_54 = x_134; -x_55 = x_136; -x_56 = x_137; -x_57 = x_140; +x_54 = x_136; +x_55 = x_137; +x_56 = x_138; +x_57 = x_139; x_58 = x_145; -x_59 = x_138; +x_59 = x_135; x_60 = x_146; goto block_133; } block_166: { -if (lean_obj_tag(x_149) == 0) +if (lean_obj_tag(x_150) == 0) { lean_dec_ref(x_153); -lean_inc_ref(x_150); -x_54 = x_154; -x_55 = x_150; +lean_inc_ref(x_148); +x_54 = x_148; +x_55 = x_149; x_56 = x_151; -x_57 = x_152; -x_58 = x_150; +x_57 = x_154; +x_58 = x_148; x_59 = x_155; x_60 = x_156; goto block_133; @@ -11657,25 +11657,25 @@ goto block_133; else { lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; -x_157 = lean_ctor_get(x_149, 0); +x_157 = lean_ctor_get(x_150, 0); lean_inc(x_157); -lean_dec_ref(x_149); +lean_dec_ref(x_150); x_158 = l_Lean_Syntax_TSepArray_getElems___redArg(x_157); lean_dec(x_157); x_159 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1___closed__18; x_160 = lean_array_get_size(x_158); -x_161 = lean_nat_dec_lt(x_148, x_160); +x_161 = lean_nat_dec_lt(x_152, x_160); if (x_161 == 0) { lean_dec(x_160); lean_dec_ref(x_158); -x_134 = x_154; -x_135 = x_153; -x_136 = x_150; -x_137 = x_151; -x_138 = x_155; -x_139 = x_156; -x_140 = x_152; +x_134 = x_156; +x_135 = x_155; +x_136 = x_148; +x_137 = x_149; +x_138 = x_151; +x_139 = x_154; +x_140 = x_153; x_141 = x_159; goto block_147; } @@ -11687,13 +11687,13 @@ if (x_162 == 0) { lean_dec(x_160); lean_dec_ref(x_158); -x_134 = x_154; -x_135 = x_153; -x_136 = x_150; -x_137 = x_151; -x_138 = x_155; -x_139 = x_156; -x_140 = x_152; +x_134 = x_156; +x_135 = x_155; +x_136 = x_148; +x_137 = x_149; +x_138 = x_151; +x_139 = x_154; +x_140 = x_153; x_141 = x_159; goto block_147; } @@ -11705,13 +11705,13 @@ x_164 = lean_usize_of_nat(x_160); lean_dec(x_160); x_165 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__9(x_158, x_163, x_164, x_159); lean_dec_ref(x_158); -x_134 = x_154; -x_135 = x_153; -x_136 = x_150; -x_137 = x_151; -x_138 = x_155; -x_139 = x_156; -x_140 = x_152; +x_134 = x_156; +x_135 = x_155; +x_136 = x_148; +x_137 = x_149; +x_138 = x_151; +x_139 = x_154; +x_140 = x_153; x_141 = x_165; goto block_147; } @@ -11722,12 +11722,12 @@ goto block_147; { lean_object* x_177; size_t x_178; size_t x_179; lean_object* x_180; x_177 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_177, 0, x_174); -lean_ctor_set(x_177, 1, x_169); +lean_ctor_set(x_177, 0, x_172); +lean_ctor_set(x_177, 1, x_171); x_178 = lean_array_size(x_176); x_179 = 0; -lean_inc_ref(x_173); -x_180 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__12(x_176, x_178, x_179, x_177, x_173, x_170); +lean_inc_ref(x_175); +x_180 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__12(x_176, x_178, x_179, x_177, x_175, x_174); lean_dec_ref(x_176); if (lean_obj_tag(x_180) == 0) { @@ -11742,25 +11742,25 @@ lean_inc(x_183); x_184 = lean_ctor_get(x_181, 1); lean_inc(x_184); lean_dec(x_181); -x_148 = x_168; -x_149 = x_167; -x_150 = x_171; -x_151 = x_172; -x_152 = x_175; +x_148 = x_167; +x_149 = x_168; +x_150 = x_170; +x_151 = x_169; +x_152 = x_173; x_153 = x_183; x_154 = x_184; -x_155 = x_173; +x_155 = x_175; x_156 = x_182; goto block_166; } else { uint8_t x_185; -lean_dec(x_175); -lean_dec_ref(x_173); -lean_dec(x_172); -lean_dec_ref(x_171); -lean_dec(x_167); +lean_dec_ref(x_175); +lean_dec(x_170); +lean_dec(x_169); +lean_dec(x_168); +lean_dec_ref(x_167); x_185 = !lean_is_exclusive(x_180); if (x_185 == 0) { @@ -11787,10 +11787,10 @@ uint8_t x_199; x_199 = l_Array_isEmpty___redArg(x_196); if (x_199 == 0) { -x_148 = x_191; -x_149 = x_190; -x_150 = x_192; -x_151 = x_193; +x_148 = x_190; +x_149 = x_191; +x_150 = x_193; +x_151 = x_192; x_152 = x_194; x_153 = x_195; x_154 = x_196; @@ -11809,20 +11809,20 @@ lean_inc_ref(x_202); x_203 = lean_mk_empty_array_with_capacity(x_201); lean_dec(x_201); x_204 = lean_array_get_size(x_202); -x_205 = lean_nat_dec_lt(x_191, x_204); +x_205 = lean_nat_dec_lt(x_194, x_204); if (x_205 == 0) { lean_dec(x_204); lean_dec_ref(x_202); x_167 = x_190; x_168 = x_191; -x_169 = x_196; -x_170 = x_198; -x_171 = x_192; -x_172 = x_193; -x_173 = x_197; -x_174 = x_195; -x_175 = x_194; +x_169 = x_192; +x_170 = x_193; +x_171 = x_196; +x_172 = x_195; +x_173 = x_194; +x_174 = x_198; +x_175 = x_197; x_176 = x_203; goto block_189; } @@ -11836,13 +11836,13 @@ lean_dec(x_204); lean_dec_ref(x_202); x_167 = x_190; x_168 = x_191; -x_169 = x_196; -x_170 = x_198; -x_171 = x_192; -x_172 = x_193; -x_173 = x_197; -x_174 = x_195; -x_175 = x_194; +x_169 = x_192; +x_170 = x_193; +x_171 = x_196; +x_172 = x_195; +x_173 = x_194; +x_174 = x_198; +x_175 = x_197; x_176 = x_203; goto block_189; } @@ -11856,13 +11856,13 @@ x_209 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00St lean_dec_ref(x_202); x_167 = x_190; x_168 = x_191; -x_169 = x_196; -x_170 = x_198; -x_171 = x_192; -x_172 = x_193; -x_173 = x_197; -x_174 = x_195; -x_175 = x_194; +x_169 = x_192; +x_170 = x_193; +x_171 = x_196; +x_172 = x_195; +x_173 = x_194; +x_174 = x_198; +x_175 = x_197; x_176 = x_209; goto block_189; } @@ -11877,11 +11877,11 @@ x_216 = lean_unsigned_to_nat(0u); x_217 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__9_spec__9___closed__0; if (lean_obj_tag(x_211) == 0) { -x_190 = x_211; -x_191 = x_216; -x_192 = x_217; -x_193 = x_212; -x_194 = x_215; +x_190 = x_217; +x_191 = x_215; +x_192 = x_212; +x_193 = x_211; +x_194 = x_216; x_195 = x_217; x_196 = x_217; x_197 = x_213; @@ -11912,11 +11912,11 @@ lean_inc(x_226); x_227 = lean_ctor_get(x_224, 1); lean_inc(x_227); lean_dec(x_224); -x_190 = x_211; -x_191 = x_216; -x_192 = x_217; -x_193 = x_212; -x_194 = x_215; +x_190 = x_217; +x_191 = x_215; +x_192 = x_212; +x_193 = x_211; +x_194 = x_216; x_195 = x_226; x_196 = x_227; x_197 = x_213; @@ -12005,53 +12005,53 @@ goto block_232; { size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_22 = lean_array_size(x_21); -lean_inc(x_15); -lean_inc_ref(x_10); -lean_inc_ref(x_13); -lean_inc_ref(x_12); -x_23 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__5(x_12, x_13, x_10, x_15, x_22, x_20, x_21); +lean_inc(x_14); +lean_inc_ref(x_17); +lean_inc_ref(x_6); +lean_inc_ref(x_5); +x_23 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1_spec__5(x_5, x_6, x_17, x_14, x_22, x_11, x_21); x_24 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1___closed__0; x_25 = l_Lean_mkSepArray(x_23, x_24); lean_dec_ref(x_23); -x_26 = l_Array_append___redArg(x_19, x_25); +x_26 = l_Array_append___redArg(x_18, x_25); lean_dec_ref(x_25); -lean_inc(x_18); -lean_inc(x_15); +lean_inc(x_9); +lean_inc(x_14); x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_18); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_9); lean_ctor_set(x_27, 2, x_26); -lean_inc(x_11); -lean_inc(x_15); -x_28 = l_Lean_Syntax_node1(x_15, x_11, x_27); -lean_inc(x_5); -lean_inc(x_15); -x_29 = l_Lean_Syntax_node1(x_15, x_5, x_28); -lean_inc(x_15); -x_30 = l_Lean_Syntax_node3(x_15, x_17, x_7, x_29, x_16); +lean_inc(x_12); +lean_inc(x_14); +x_28 = l_Lean_Syntax_node1(x_14, x_12, x_27); +lean_inc(x_10); +lean_inc(x_14); +x_29 = l_Lean_Syntax_node1(x_14, x_10, x_28); +lean_inc(x_14); +x_30 = l_Lean_Syntax_node3(x_14, x_20, x_16, x_29, x_19); x_31 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticWf__trivial__1___closed__147; -lean_inc(x_15); +lean_inc(x_14); x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_15); +lean_ctor_set(x_32, 0, x_14); lean_ctor_set(x_32, 1, x_31); x_33 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1___closed__1; -x_34 = l_Lean_Name_mkStr4(x_12, x_13, x_10, x_33); +x_34 = l_Lean_Name_mkStr4(x_5, x_6, x_17, x_33); x_35 = l_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__model_x5b___x5dUsing____1___closed__2; -lean_inc(x_15); +lean_inc(x_14); x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_15); +lean_ctor_set(x_36, 0, x_14); lean_ctor_set(x_36, 1, x_35); -lean_inc(x_15); -x_37 = l_Lean_Syntax_node2(x_15, x_34, x_36, x_4); -lean_inc(x_15); -x_38 = l_Lean_Syntax_node1(x_15, x_18, x_37); -lean_inc(x_15); -x_39 = l_Lean_Syntax_node1(x_15, x_11, x_38); -lean_inc(x_15); -x_40 = l_Lean_Syntax_node1(x_15, x_5, x_39); -lean_inc(x_15); -x_41 = l_Lean_Syntax_node2(x_15, x_6, x_9, x_40); -x_42 = l_Lean_Syntax_node3(x_15, x_14, x_30, x_32, x_41); +lean_inc(x_14); +x_37 = l_Lean_Syntax_node2(x_14, x_34, x_36, x_4); +lean_inc(x_14); +x_38 = l_Lean_Syntax_node1(x_14, x_9, x_37); +lean_inc(x_14); +x_39 = l_Lean_Syntax_node1(x_14, x_12, x_38); +lean_inc(x_14); +x_40 = l_Lean_Syntax_node1(x_14, x_10, x_39); +lean_inc(x_14); +x_41 = l_Lean_Syntax_node2(x_14, x_13, x_7, x_40); +x_42 = l_Lean_Syntax_node3(x_14, x_15, x_30, x_32, 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_8); diff --git a/stage0/stdlib/Std/Data/DTreeMap/Internal/Lemmas.c b/stage0/stdlib/Std/Data/DTreeMap/Internal/Lemmas.c index bf68b685e854..f51eb49707a3 100644 --- a/stage0/stdlib/Std/Data/DTreeMap/Internal/Lemmas.c +++ b/stage0/stdlib/Std/Data/DTreeMap/Internal/Lemmas.c @@ -13771,7 +13771,7 @@ return x_2; LEAN_EXPORT lean_object* l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; x_45 = l_Std_DTreeMap_Internal_Impl_tacticWf__trivial___closed__0; x_46 = l_Std_DTreeMap_Internal_Impl_tacticWf__trivial___closed__1; x_47 = l_Std_DTreeMap_Internal_Impl_tacticWf__trivial___closed__2; @@ -13842,7 +13842,7 @@ goto block_247; block_135: { lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; size_t x_115; size_t x_116; lean_object* x_117; size_t x_118; lean_object* x_119; lean_object* x_120; size_t 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_dec(x_54); +lean_dec(x_55); x_61 = lean_ctor_get(x_59, 5); lean_inc(x_61); lean_dec_ref(x_59); @@ -13949,8 +13949,8 @@ x_112 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_112, 0, x_63); lean_ctor_set(x_112, 1, x_111); x_113 = l___private_Std_Data_DTreeMap_Internal_Lemmas_0__Std_DTreeMap_Internal_Impl_helperLemmaNames; -x_114 = l_Array_append___redArg(x_113, x_57); -lean_dec_ref(x_57); +x_114 = l_Array_append___redArg(x_113, x_56); +lean_dec_ref(x_56); x_115 = lean_array_size(x_114); x_116 = 0; x_117 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__2(x_115, x_116, x_114); @@ -13983,54 +13983,54 @@ lean_inc_ref(x_79); lean_inc(x_63); x_130 = l_Lean_Syntax_node6(x_63, x_75, x_76, x_80, x_107, x_110, x_129, x_79); x_131 = l_Array_mkArray2___redArg(x_130, x_79); -if (lean_obj_tag(x_56) == 0) -{ -x_4 = x_90; -x_5 = x_105; -x_6 = x_66; -x_7 = x_99; -x_8 = x_67; -x_9 = x_60; -x_10 = x_65; -x_11 = x_116; +if (lean_obj_tag(x_54) == 0) +{ +x_4 = x_70; +x_5 = x_92; +x_6 = x_67; +x_7 = x_71; +x_8 = x_64; +x_9 = x_65; +x_10 = x_99; +x_11 = x_63; x_12 = x_72; -x_13 = x_92; -x_14 = x_63; -x_15 = x_68; -x_16 = x_131; -x_17 = x_64; -x_18 = x_70; -x_19 = x_73; -x_20 = x_71; -x_21 = x_55; +x_13 = x_105; +x_14 = x_131; +x_15 = x_73; +x_16 = x_68; +x_17 = x_60; +x_18 = x_90; +x_19 = x_66; +x_20 = x_116; +x_21 = x_57; goto block_44; } else { lean_object* x_132; lean_object* x_133; lean_object* x_134; -lean_dec_ref(x_55); -x_132 = lean_ctor_get(x_56, 0); +lean_dec_ref(x_57); +x_132 = lean_ctor_get(x_54, 0); lean_inc(x_132); -lean_dec_ref(x_56); +lean_dec_ref(x_54); x_133 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1___closed__17; x_134 = lean_array_push(x_133, x_132); -x_4 = x_90; -x_5 = x_105; -x_6 = x_66; -x_7 = x_99; -x_8 = x_67; -x_9 = x_60; -x_10 = x_65; -x_11 = x_116; +x_4 = x_70; +x_5 = x_92; +x_6 = x_67; +x_7 = x_71; +x_8 = x_64; +x_9 = x_65; +x_10 = x_99; +x_11 = x_63; x_12 = x_72; -x_13 = x_92; -x_14 = x_63; -x_15 = x_68; -x_16 = x_131; -x_17 = x_64; -x_18 = x_70; -x_19 = x_73; -x_20 = x_71; +x_13 = x_105; +x_14 = x_131; +x_15 = x_73; +x_16 = x_68; +x_17 = x_60; +x_18 = x_90; +x_19 = x_66; +x_20 = x_116; x_21 = x_134; goto block_44; } @@ -14040,36 +14040,36 @@ goto block_44; size_t x_144; size_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; x_144 = lean_array_size(x_143); x_145 = 0; +lean_inc_ref(x_139); lean_inc_ref(x_142); -lean_inc_ref(x_137); -x_146 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__7(x_136, x_45, x_46, x_47, x_48, x_139, x_143, x_144, x_145, x_137, x_142, x_138); +x_146 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__7(x_138, x_45, x_46, x_47, x_48, x_140, x_143, x_144, x_145, x_142, x_139, x_141); lean_dec_ref(x_143); -lean_dec_ref(x_139); +lean_dec_ref(x_140); 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_54 = x_136; -x_55 = x_137; -x_56 = x_141; -x_57 = x_140; +x_55 = x_138; +x_56 = x_137; +x_57 = x_142; x_58 = x_147; -x_59 = x_142; +x_59 = x_139; x_60 = x_148; goto block_135; } block_168: { -if (lean_obj_tag(x_152) == 0) +if (lean_obj_tag(x_154) == 0) { lean_dec_ref(x_155); -lean_inc_ref(x_151); +lean_inc_ref(x_153); x_54 = x_150; x_55 = x_151; -x_56 = x_153; -x_57 = x_156; -x_58 = x_151; +x_56 = x_156; +x_57 = x_153; +x_58 = x_153; x_59 = x_157; x_60 = x_158; goto block_135; @@ -14077,25 +14077,25 @@ goto block_135; else { lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; -x_159 = lean_ctor_get(x_152, 0); +x_159 = lean_ctor_get(x_154, 0); lean_inc(x_159); -lean_dec_ref(x_152); +lean_dec_ref(x_154); x_160 = l_Lean_Syntax_TSepArray_getElems___redArg(x_159); lean_dec(x_159); x_161 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1___closed__18; x_162 = lean_array_get_size(x_160); -x_163 = lean_nat_dec_lt(x_154, x_162); +x_163 = lean_nat_dec_lt(x_152, x_162); if (x_163 == 0) { lean_dec(x_162); lean_dec_ref(x_160); x_136 = x_150; -x_137 = x_151; -x_138 = x_158; -x_139 = x_155; -x_140 = x_156; -x_141 = x_153; -x_142 = x_157; +x_137 = x_156; +x_138 = x_151; +x_139 = x_157; +x_140 = x_155; +x_141 = x_158; +x_142 = x_153; x_143 = x_161; goto block_149; } @@ -14108,12 +14108,12 @@ if (x_164 == 0) lean_dec(x_162); lean_dec_ref(x_160); x_136 = x_150; -x_137 = x_151; -x_138 = x_158; -x_139 = x_155; -x_140 = x_156; -x_141 = x_153; -x_142 = x_157; +x_137 = x_156; +x_138 = x_151; +x_139 = x_157; +x_140 = x_155; +x_141 = x_158; +x_142 = x_153; x_143 = x_161; goto block_149; } @@ -14126,12 +14126,12 @@ lean_dec(x_162); x_167 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__9(x_160, x_165, x_166, x_161); lean_dec_ref(x_160); x_136 = x_150; -x_137 = x_151; -x_138 = x_158; -x_139 = x_155; -x_140 = x_156; -x_141 = x_153; -x_142 = x_157; +x_137 = x_156; +x_138 = x_151; +x_139 = x_157; +x_140 = x_155; +x_141 = x_158; +x_142 = x_153; x_143 = x_167; goto block_149; } @@ -14142,12 +14142,12 @@ goto block_149; { lean_object* x_179; size_t x_180; size_t x_181; lean_object* x_182; x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_171); -lean_ctor_set(x_179, 1, x_173); +lean_ctor_set(x_179, 0, x_174); +lean_ctor_set(x_179, 1, x_175); x_180 = lean_array_size(x_178); x_181 = 0; -lean_inc_ref(x_177); -x_182 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__12(x_178, x_180, x_181, x_179, x_177, x_176); +lean_inc_ref(x_173); +x_182 = l___private_Init_Data_Array_Basic_0__Array_forIn_x27Unsafe_loop___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__12(x_178, x_180, x_181, x_179, x_173, x_172); lean_dec_ref(x_178); if (lean_obj_tag(x_182) == 0) { @@ -14164,22 +14164,22 @@ lean_inc(x_186); lean_dec(x_183); x_150 = x_169; x_151 = x_170; -x_152 = x_172; -x_153 = x_174; -x_154 = x_175; +x_152 = x_171; +x_153 = x_176; +x_154 = x_177; x_155 = x_185; x_156 = x_186; -x_157 = x_177; +x_157 = x_173; x_158 = x_184; goto block_168; } else { uint8_t x_187; -lean_dec_ref(x_177); -lean_dec(x_174); -lean_dec(x_172); -lean_dec_ref(x_170); +lean_dec(x_177); +lean_dec_ref(x_176); +lean_dec_ref(x_173); +lean_dec(x_170); lean_dec(x_169); x_187 = !lean_is_exclusive(x_182); if (x_187 == 0) @@ -14229,20 +14229,20 @@ lean_inc_ref(x_204); x_205 = lean_mk_empty_array_with_capacity(x_203); lean_dec(x_203); x_206 = lean_array_get_size(x_204); -x_207 = lean_nat_dec_lt(x_196, x_206); +x_207 = lean_nat_dec_lt(x_194, x_206); if (x_207 == 0) { lean_dec(x_206); lean_dec_ref(x_204); x_169 = x_192; x_170 = x_193; -x_171 = x_197; -x_172 = x_194; -x_173 = x_198; -x_174 = x_195; -x_175 = x_196; -x_176 = x_200; -x_177 = x_199; +x_171 = x_194; +x_172 = x_200; +x_173 = x_199; +x_174 = x_197; +x_175 = x_198; +x_176 = x_195; +x_177 = x_196; x_178 = x_205; goto block_191; } @@ -14256,13 +14256,13 @@ lean_dec(x_206); lean_dec_ref(x_204); x_169 = x_192; x_170 = x_193; -x_171 = x_197; -x_172 = x_194; -x_173 = x_198; -x_174 = x_195; -x_175 = x_196; -x_176 = x_200; -x_177 = x_199; +x_171 = x_194; +x_172 = x_200; +x_173 = x_199; +x_174 = x_197; +x_175 = x_198; +x_176 = x_195; +x_177 = x_196; x_178 = x_205; goto block_191; } @@ -14276,13 +14276,13 @@ x_211 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00St lean_dec_ref(x_204); x_169 = x_192; x_170 = x_193; -x_171 = x_197; -x_172 = x_194; -x_173 = x_198; -x_174 = x_195; -x_175 = x_196; -x_176 = x_200; -x_177 = x_199; +x_171 = x_194; +x_172 = x_200; +x_173 = x_199; +x_174 = x_197; +x_175 = x_198; +x_176 = x_195; +x_177 = x_196; x_178 = x_211; goto block_191; } @@ -14297,11 +14297,11 @@ x_218 = lean_unsigned_to_nat(0u); x_219 = l___private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00__private_Init_Data_Array_Basic_0__Array_foldlMUnsafe_fold___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__9_spec__9___closed__0; if (lean_obj_tag(x_213) == 0) { -x_192 = x_217; -x_193 = x_219; -x_194 = x_213; -x_195 = x_214; -x_196 = x_218; +x_192 = x_214; +x_193 = x_217; +x_194 = x_218; +x_195 = x_219; +x_196 = x_213; x_197 = x_219; x_198 = x_219; x_199 = x_215; @@ -14332,11 +14332,11 @@ lean_inc(x_228); x_229 = lean_ctor_get(x_226, 1); lean_inc(x_229); lean_dec(x_226); -x_192 = x_217; -x_193 = x_219; -x_194 = x_213; -x_195 = x_214; -x_196 = x_218; +x_192 = x_214; +x_193 = x_217; +x_194 = x_218; +x_195 = x_219; +x_196 = x_213; x_197 = x_228; x_198 = x_229; x_199 = x_215; @@ -14425,56 +14425,56 @@ goto block_234; { size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_22 = lean_array_size(x_21); -lean_inc(x_14); -lean_inc_ref(x_6); -lean_inc_ref(x_10); -lean_inc_ref(x_17); -x_23 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__5(x_17, x_10, x_6, x_14, x_22, x_11, x_21); +lean_inc(x_11); +lean_inc_ref(x_19); +lean_inc_ref(x_9); +lean_inc_ref(x_8); +x_23 = l___private_Init_Data_Array_Basic_0__Array_mapMUnsafe_map___at___00Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1_spec__5(x_8, x_9, x_19, x_11, x_22, x_20, x_21); x_24 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1___closed__0; x_25 = l_Lean_mkSepArray(x_23, x_24); lean_dec_ref(x_23); -x_26 = l_Array_append___redArg(x_16, x_25); +x_26 = l_Array_append___redArg(x_14, x_25); lean_dec_ref(x_25); -lean_inc(x_19); -lean_inc(x_14); +lean_inc(x_15); +lean_inc(x_11); x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 0, x_11); +lean_ctor_set(x_27, 1, x_15); lean_ctor_set(x_27, 2, x_26); lean_inc(x_12); -lean_inc(x_14); -x_28 = l_Lean_Syntax_node1(x_14, x_12, x_27); -lean_inc(x_20); -lean_inc(x_14); -x_29 = l_Lean_Syntax_node1(x_14, x_20, x_28); -lean_inc(x_14); -x_30 = l_Lean_Syntax_node3(x_14, x_15, x_18, x_29, x_5); +lean_inc(x_11); +x_28 = l_Lean_Syntax_node1(x_11, x_12, x_27); +lean_inc(x_7); +lean_inc(x_11); +x_29 = l_Lean_Syntax_node1(x_11, x_7, x_28); +lean_inc(x_11); +x_30 = l_Lean_Syntax_node3(x_11, x_16, x_4, x_29, x_13); x_31 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticWf__trivial__1___closed__146; -lean_inc(x_14); +lean_inc(x_11); x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_14); +lean_ctor_set(x_32, 0, x_11); lean_ctor_set(x_32, 1, x_31); x_33 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1___closed__1; -x_34 = l_Lean_Name_mkStr4(x_17, x_10, x_6, x_33); +x_34 = l_Lean_Name_mkStr4(x_8, x_9, x_19, x_33); x_35 = l_Std_DTreeMap_Internal_Impl___aux__Std__Data__DTreeMap__Internal__Lemmas______macroRules__Std__DTreeMap__Internal__Impl__tacticSimp__to__model_x5b___x5dUsing____1___closed__2; -lean_inc(x_14); +lean_inc(x_11); x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_14); +lean_ctor_set(x_36, 0, x_11); lean_ctor_set(x_36, 1, x_35); -lean_inc(x_14); -x_37 = l_Lean_Syntax_node2(x_14, x_34, x_36, x_7); -lean_inc(x_14); -x_38 = l_Lean_Syntax_node1(x_14, x_19, x_37); -lean_inc(x_14); -x_39 = l_Lean_Syntax_node1(x_14, x_12, x_38); -lean_inc(x_14); -x_40 = l_Lean_Syntax_node1(x_14, x_20, x_39); -lean_inc(x_14); -x_41 = l_Lean_Syntax_node2(x_14, x_4, x_13, x_40); -x_42 = l_Lean_Syntax_node3(x_14, x_8, x_30, x_32, x_41); +lean_inc(x_11); +x_37 = l_Lean_Syntax_node2(x_11, x_34, x_36, x_10); +lean_inc(x_11); +x_38 = l_Lean_Syntax_node1(x_11, x_15, x_37); +lean_inc(x_11); +x_39 = l_Lean_Syntax_node1(x_11, x_12, x_38); +lean_inc(x_11); +x_40 = l_Lean_Syntax_node1(x_11, x_7, x_39); +lean_inc(x_11); +x_41 = l_Lean_Syntax_node2(x_11, x_18, x_5, x_40); +x_42 = l_Lean_Syntax_node3(x_11, x_6, x_30, x_32, 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_9); +lean_ctor_set(x_43, 1, x_17); return x_43; } } diff --git a/tests/lean/sint_basic.lean.expected.out b/tests/lean/sint_basic.lean.expected.out index e605b11dd600..183e56a61ac4 100644 --- a/tests/lean/sint_basic.lean.expected.out +++ b/tests/lean/sint_basic.lean.expected.out @@ -73,11 +73,12 @@ true true true [Compiler.IR] [result] - def _private.lean.sint_basic.0.myId8 (x_1 : u8) : u8 := + def _private.lean.sint_basic.0.myId8 (x_1 : i8) : i8 := ret x_1 - def _private.lean.sint_basic.0.myId8._boxed (x_1 : tagged) : tagged := - let x_2 : u8 := unbox x_1; - let x_3 : u8 := _private.lean.sint_basic.0.myId8 x_2; + def _private.lean.sint_basic.0.myId8._boxed (x_1 : tobj) : tobj := + let x_2 : i8 := unbox x_1; + dec x_1; + let x_3 : i8 := _private.lean.sint_basic.0.myId8 x_2; let x_4 : tobj := box x_3; ret x_4 Int16 : Type @@ -155,11 +156,12 @@ true true true [Compiler.IR] [result] - def _private.lean.sint_basic.0.myId16 (x_1 : u16) : u16 := + def _private.lean.sint_basic.0.myId16 (x_1 : i16) : i16 := ret x_1 - def _private.lean.sint_basic.0.myId16._boxed (x_1 : tagged) : tagged := - let x_2 : u16 := unbox x_1; - let x_3 : u16 := _private.lean.sint_basic.0.myId16 x_2; + def _private.lean.sint_basic.0.myId16._boxed (x_1 : tobj) : tobj := + let x_2 : i16 := unbox x_1; + dec x_1; + let x_3 : i16 := _private.lean.sint_basic.0.myId16 x_2; let x_4 : tobj := box x_3; ret x_4 Int32 : Type @@ -237,12 +239,12 @@ true true true [Compiler.IR] [result] - def _private.lean.sint_basic.0.myId32 (x_1 : u32) : u32 := + def _private.lean.sint_basic.0.myId32 (x_1 : i32) : i32 := ret x_1 def _private.lean.sint_basic.0.myId32._boxed (x_1 : tobj) : tobj := - let x_2 : u32 := unbox x_1; + let x_2 : i32 := unbox x_1; dec x_1; - let x_3 : u32 := _private.lean.sint_basic.0.myId32 x_2; + let x_3 : i32 := _private.lean.sint_basic.0.myId32 x_2; let x_4 : tobj := box x_3; ret x_4 Int64 : Type @@ -320,12 +322,12 @@ true true true [Compiler.IR] [result] - def _private.lean.sint_basic.0.myId64 (x_1 : u64) : u64 := + def _private.lean.sint_basic.0.myId64 (x_1 : i64) : i64 := ret x_1 def _private.lean.sint_basic.0.myId64._boxed (x_1 : tobj) : tobj := - let x_2 : u64 := unbox x_1; + let x_2 : i64 := unbox x_1; dec x_1; - let x_3 : u64 := _private.lean.sint_basic.0.myId64 x_2; + let x_3 : i64 := _private.lean.sint_basic.0.myId64 x_2; let x_4 : tobj := box x_3; ret x_4 ISize : Type @@ -403,11 +405,11 @@ true true true [Compiler.IR] [result] - def _private.lean.sint_basic.0.myIdSize (x_1 : usize) : usize := + def _private.lean.sint_basic.0.myIdSize (x_1 : isize) : isize := ret x_1 def _private.lean.sint_basic.0.myIdSize._boxed (x_1 : tobj) : tobj := - let x_2 : usize := unbox x_1; + let x_2 : isize := unbox x_1; dec x_1; - let x_3 : usize := _private.lean.sint_basic.0.myIdSize x_2; + let x_3 : isize := _private.lean.sint_basic.0.myIdSize x_2; let x_4 : tobj := box x_3; ret x_4