diff --git a/stage0/src/CMakeLists.txt b/stage0/src/CMakeLists.txt index 06991773883e..2b56508725a4 100644 --- a/stage0/src/CMakeLists.txt +++ b/stage0/src/CMakeLists.txt @@ -12,7 +12,7 @@ project(LEAN CXX C) set(LEAN_VERSION_MAJOR 4) set(LEAN_VERSION_MINOR 11) set(LEAN_VERSION_PATCH 0) -set(LEAN_VERSION_IS_RELEASE 0) # This number is 1 in the release revision, and 0 otherwise. +set(LEAN_VERSION_IS_RELEASE 1) # This number is 1 in the release revision, and 0 otherwise. set(LEAN_SPECIAL_VERSION_DESC "" CACHE STRING "Additional version description like 'nightly-2018-03-11'") set(LEAN_VERSION_STRING "${LEAN_VERSION_MAJOR}.${LEAN_VERSION_MINOR}.${LEAN_VERSION_PATCH}") if (LEAN_SPECIAL_VERSION_DESC) diff --git a/stage0/src/kernel/declaration.h b/stage0/src/kernel/declaration.h index cf45393b435c..9b636ae0a8f6 100644 --- a/stage0/src/kernel/declaration.h +++ b/stage0/src/kernel/declaration.h @@ -63,9 +63,9 @@ class constant_val : public object_ref { public: constant_val(name const & n, names const & lparams, expr const & type); constant_val(constant_val const & other):object_ref(other) {} - constant_val(constant_val && other):object_ref(other) {} + constant_val(constant_val && other):object_ref(std::move(other)) {} constant_val & operator=(constant_val const & other) { object_ref::operator=(other); return *this; } - constant_val & operator=(constant_val && other) { object_ref::operator=(other); return *this; } + constant_val & operator=(constant_val && other) { object_ref::operator=(std::move(other)); return *this; } name const & get_name() const { return static_cast(cnstr_get_ref(*this, 0)); } names const & get_lparams() const { return static_cast(cnstr_get_ref(*this, 1)); } expr const & get_type() const { return static_cast(cnstr_get_ref(*this, 2)); } @@ -79,9 +79,9 @@ class axiom_val : public object_ref { public: axiom_val(name const & n, names const & lparams, expr const & type, bool is_unsafe); axiom_val(axiom_val const & other):object_ref(other) {} - axiom_val(axiom_val && other):object_ref(other) {} + axiom_val(axiom_val && other):object_ref(std::move(other)) {} axiom_val & operator=(axiom_val const & other) { object_ref::operator=(other); return *this; } - axiom_val & operator=(axiom_val && other) { object_ref::operator=(other); return *this; } + axiom_val & operator=(axiom_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } names const & get_lparams() const { return to_constant_val().get_lparams(); } @@ -105,9 +105,9 @@ class definition_val : public object_ref { public: definition_val(name const & n, names const & lparams, expr const & type, expr const & val, reducibility_hints const & hints, definition_safety safety, names const & all); definition_val(definition_val const & other):object_ref(other) {} - definition_val(definition_val && other):object_ref(other) {} + definition_val(definition_val && other):object_ref(std::move(other)) {} definition_val & operator=(definition_val const & other) { object_ref::operator=(other); return *this; } - definition_val & operator=(definition_val && other) { object_ref::operator=(other); return *this; } + definition_val & operator=(definition_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } names const & get_lparams() const { return to_constant_val().get_lparams(); } @@ -128,9 +128,9 @@ class theorem_val : public object_ref { public: theorem_val(name const & n, names const & lparams, expr const & type, expr const & val, names const & all); theorem_val(theorem_val const & other):object_ref(other) {} - theorem_val(theorem_val && other):object_ref(other) {} + theorem_val(theorem_val && other):object_ref(std::move(other)) {} theorem_val & operator=(theorem_val const & other) { object_ref::operator=(other); return *this; } - theorem_val & operator=(theorem_val && other) { object_ref::operator=(other); return *this; } + theorem_val & operator=(theorem_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } names const & get_lparams() const { return to_constant_val().get_lparams(); } @@ -148,9 +148,9 @@ class opaque_val : public object_ref { public: opaque_val(name const & n, names const & lparams, expr const & type, expr const & val, bool is_unsafe, names const & all); opaque_val(opaque_val const & other):object_ref(other) {} - opaque_val(opaque_val && other):object_ref(other) {} + opaque_val(opaque_val && other):object_ref(std::move(other)) {} opaque_val & operator=(opaque_val const & other) { object_ref::operator=(other); return *this; } - opaque_val & operator=(opaque_val && other) { object_ref::operator=(other); return *this; } + opaque_val & operator=(opaque_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } names const & get_lparams() const { return to_constant_val().get_lparams(); } @@ -179,9 +179,9 @@ class inductive_type : public object_ref { public: inductive_type(name const & id, expr const & type, constructors const & cnstrs); inductive_type(inductive_type const & other):object_ref(other) {} - inductive_type(inductive_type && other):object_ref(other) {} + inductive_type(inductive_type && other):object_ref(std::move(other)) {} inductive_type & operator=(inductive_type const & other) { object_ref::operator=(other); return *this; } - inductive_type & operator=(inductive_type && other) { object_ref::operator=(other); return *this; } + inductive_type & operator=(inductive_type && other) { object_ref::operator=(std::move(other)); return *this; } name const & get_name() const { return static_cast(cnstr_get_ref(*this, 0)); } expr const & get_type() const { return static_cast(cnstr_get_ref(*this, 1)); } constructors const & get_cnstrs() const { return static_cast(cnstr_get_ref(*this, 2)); } @@ -205,7 +205,7 @@ class declaration : public object_ref { public: declaration(); declaration(declaration const & other):object_ref(other) {} - declaration(declaration && other):object_ref(other) {} + declaration(declaration && other):object_ref(std::move(other)) {} /* low-level constructors */ explicit declaration(object * o):object_ref(o) {} explicit declaration(b_obj_arg o, bool b):object_ref(o, b) {} @@ -213,7 +213,7 @@ class declaration : public object_ref { declaration_kind kind() const { return static_cast(obj_tag(raw())); } declaration & operator=(declaration const & other) { object_ref::operator=(other); return *this; } - declaration & operator=(declaration && other) { object_ref::operator=(other); return *this; } + declaration & operator=(declaration && other) { object_ref::operator=(std::move(other)); return *this; } friend bool is_eqp(declaration const & d1, declaration const & d2) { return d1.raw() == d2.raw(); } @@ -263,10 +263,10 @@ declaration mk_axiom_inferring_unsafe(environment const & env, name const & n, class inductive_decl : public object_ref { public: inductive_decl(inductive_decl const & other):object_ref(other) {} - inductive_decl(inductive_decl && other):object_ref(other) {} + inductive_decl(inductive_decl && other):object_ref(std::move(other)) {} inductive_decl(declaration const & d):object_ref(d) { lean_assert(d.is_inductive()); } inductive_decl & operator=(inductive_decl const & other) { object_ref::operator=(other); return *this; } - inductive_decl & operator=(inductive_decl && other) { object_ref::operator=(other); return *this; } + inductive_decl & operator=(inductive_decl && other) { object_ref::operator=(std::move(other)); return *this; } names const & get_lparams() const { return static_cast(cnstr_get_ref(raw(), 0)); } nat const & get_nparams() const { return static_cast(cnstr_get_ref(raw(), 1)); } inductive_types const & get_types() const { return static_cast(cnstr_get_ref(raw(), 2)); } @@ -290,9 +290,9 @@ class inductive_val : public object_ref { inductive_val(name const & n, names const & lparams, expr const & type, unsigned nparams, unsigned nindices, names const & all, names const & cnstrs, unsigned nnested, bool is_rec, bool is_unsafe, bool is_reflexive); inductive_val(inductive_val const & other):object_ref(other) {} - inductive_val(inductive_val && other):object_ref(other) {} + inductive_val(inductive_val && other):object_ref(std::move(other)) {} inductive_val & operator=(inductive_val const & other) { object_ref::operator=(other); return *this; } - inductive_val & operator=(inductive_val && other) { object_ref::operator=(other); return *this; } + inductive_val & operator=(inductive_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } unsigned get_nparams() const { return static_cast(cnstr_get_ref(*this, 1)).get_small_value(); } unsigned get_nindices() const { return static_cast(cnstr_get_ref(*this, 2)).get_small_value(); } @@ -317,9 +317,9 @@ class constructor_val : public object_ref { public: constructor_val(name const & n, names const & lparams, expr const & type, name const & induct, unsigned cidx, unsigned nparams, unsigned nfields, bool is_unsafe); constructor_val(constructor_val const & other):object_ref(other) {} - constructor_val(constructor_val && other):object_ref(other) {} + constructor_val(constructor_val && other):object_ref(std::move(other)) {} constructor_val & operator=(constructor_val const & other) { object_ref::operator=(other); return *this; } - constructor_val & operator=(constructor_val && other) { object_ref::operator=(other); return *this; } + constructor_val & operator=(constructor_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_induct() const { return static_cast(cnstr_get_ref(*this, 1)); } unsigned get_cidx() const { return static_cast(cnstr_get_ref(*this, 2)).get_small_value(); } @@ -338,9 +338,9 @@ class recursor_rule : public object_ref { public: recursor_rule(name const & cnstr, unsigned nfields, expr const & rhs); recursor_rule(recursor_rule const & other):object_ref(other) {} - recursor_rule(recursor_rule && other):object_ref(other) {} + recursor_rule(recursor_rule && other):object_ref(std::move(other)) {} recursor_rule & operator=(recursor_rule const & other) { object_ref::operator=(other); return *this; } - recursor_rule & operator=(recursor_rule && other) { object_ref::operator=(other); return *this; } + recursor_rule & operator=(recursor_rule && other) { object_ref::operator=(std::move(other)); return *this; } name const & get_cnstr() const { return static_cast(cnstr_get_ref(*this, 0)); } unsigned get_nfields() const { return static_cast(cnstr_get_ref(*this, 1)).get_small_value(); } expr const & get_rhs() const { return static_cast(cnstr_get_ref(*this, 2)); } @@ -365,9 +365,9 @@ class recursor_val : public object_ref { names const & all, unsigned nparams, unsigned nindices, unsigned nmotives, unsigned nminors, recursor_rules const & rules, bool k, bool is_unsafe); recursor_val(recursor_val const & other):object_ref(other) {} - recursor_val(recursor_val && other):object_ref(other) {} + recursor_val(recursor_val && other):object_ref(std::move(other)) {} recursor_val & operator=(recursor_val const & other) { object_ref::operator=(other); return *this; } - recursor_val & operator=(recursor_val && other) { object_ref::operator=(other); return *this; } + recursor_val & operator=(recursor_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } name const & get_induct() const { return get_name().get_prefix(); } @@ -398,9 +398,9 @@ class quot_val : public object_ref { public: quot_val(name const & n, names const & lparams, expr const & type, quot_kind k); quot_val(quot_val const & other):object_ref(other) {} - quot_val(quot_val && other):object_ref(other) {} + quot_val(quot_val && other):object_ref(std::move(other)) {} quot_val & operator=(quot_val const & other) { object_ref::operator=(other); return *this; } - quot_val & operator=(quot_val && other) { object_ref::operator=(other); return *this; } + quot_val & operator=(quot_val && other) { object_ref::operator=(std::move(other)); return *this; } constant_val const & to_constant_val() const { return static_cast(cnstr_get_ref(*this, 0)); } name const & get_name() const { return to_constant_val().get_name(); } names const & get_lparams() const { return to_constant_val().get_lparams(); } @@ -434,14 +434,14 @@ class constant_info : public object_ref { constant_info(constructor_val const & v); constant_info(recursor_val const & v); constant_info(constant_info const & other):object_ref(other) {} - constant_info(constant_info && other):object_ref(other) {} + constant_info(constant_info && other):object_ref(std::move(other)) {} explicit constant_info(b_obj_arg o, bool b):object_ref(o, b) {} explicit constant_info(obj_arg o):object_ref(o) {} constant_info_kind kind() const { return static_cast(cnstr_tag(raw())); } constant_info & operator=(constant_info const & other) { object_ref::operator=(other); return *this; } - constant_info & operator=(constant_info && other) { object_ref::operator=(other); return *this; } + constant_info & operator=(constant_info && other) { object_ref::operator=(std::move(other)); return *this; } friend bool is_eqp(constant_info const & d1, constant_info const & d2) { return d1.raw() == d2.raw(); } diff --git a/stage0/src/kernel/expr.h b/stage0/src/kernel/expr.h index 0371110d335a..d1771e73a366 100644 --- a/stage0/src/kernel/expr.h +++ b/stage0/src/kernel/expr.h @@ -50,9 +50,9 @@ class literal : public object_ref { explicit literal(nat const & v); literal():literal(0u) {} literal(literal const & other):object_ref(other) {} - literal(literal && other):object_ref(other) {} + literal(literal && other):object_ref(std::move(other)) {} literal & operator=(literal const & other) { object_ref::operator=(other); return *this; } - literal & operator=(literal && other) { object_ref::operator=(other); return *this; } + literal & operator=(literal && other) { object_ref::operator=(std::move(other)); return *this; } static literal_kind kind(object * o) { return static_cast(cnstr_tag(o)); } literal_kind kind() const { return kind(raw()); } @@ -100,14 +100,14 @@ class expr : public object_ref { public: expr(); expr(expr const & other):object_ref(other) {} - expr(expr && other):object_ref(other) {} + expr(expr && other):object_ref(std::move(other)) {} explicit expr(b_obj_arg o, bool b):object_ref(o, b) {} explicit expr(obj_arg o):object_ref(o) {} static expr_kind kind(object * o) { return static_cast(cnstr_tag(o)); } expr_kind kind() const { return kind(raw()); } expr & operator=(expr const & other) { object_ref::operator=(other); return *this; } - expr & operator=(expr && other) { object_ref::operator=(other); return *this; } + expr & operator=(expr && other) { object_ref::operator=(std::move(other)); return *this; } friend bool is_eqp(expr const & e1, expr const & e2) { return e1.raw() == e2.raw(); } }; diff --git a/stage0/src/kernel/level.h b/stage0/src/kernel/level.h index b01240747ec9..32bfade7c2e0 100644 --- a/stage0/src/kernel/level.h +++ b/stage0/src/kernel/level.h @@ -43,14 +43,14 @@ class level : public object_ref { explicit level(obj_arg o):object_ref(o) {} explicit level(b_obj_arg o, bool b):object_ref(o, b) {} level(level const & other):object_ref(other) {} - level(level && other):object_ref(other) {} + level(level && other):object_ref(std::move(other)) {} level_kind kind() const { return lean_is_scalar(raw()) ? level_kind::Zero : static_cast(lean_ptr_tag(raw())); } unsigned hash() const; level & operator=(level const & other) { object_ref::operator=(other); return *this; } - level & operator=(level && other) { object_ref::operator=(other); return *this; } + level & operator=(level && other) { object_ref::operator=(std::move(other)); return *this; } friend bool is_eqp(level const & l1, level const & l2) { return l1.raw() == l2.raw(); } diff --git a/stage0/src/kernel/local_ctx.h b/stage0/src/kernel/local_ctx.h index 8338316a4a84..500d62a03c88 100644 --- a/stage0/src/kernel/local_ctx.h +++ b/stage0/src/kernel/local_ctx.h @@ -27,11 +27,11 @@ class local_decl : public object_ref { public: local_decl(); local_decl(local_decl const & other):object_ref(other) {} - local_decl(local_decl && other):object_ref(other) {} + local_decl(local_decl && other):object_ref(std::move(other)) {} local_decl(obj_arg o):object_ref(o) {} local_decl(b_obj_arg o, bool):object_ref(o, true) {} local_decl & operator=(local_decl const & other) { object_ref::operator=(other); return *this; } - local_decl & operator=(local_decl && other) { object_ref::operator=(other); return *this; } + local_decl & operator=(local_decl && other) { object_ref::operator=(std::move(other)); return *this; } friend bool is_eqp(local_decl const & d1, local_decl const & d2) { return d1.raw() == d2.raw(); } unsigned get_idx() const { return static_cast(cnstr_get_ref(raw(), 0)).get_small_value(); } name const & get_name() const { return static_cast(cnstr_get_ref(raw(), 1)); } @@ -54,9 +54,9 @@ class local_ctx : public object_ref { explicit local_ctx(obj_arg o):object_ref(o) {} local_ctx(b_obj_arg o, bool):object_ref(o, true) {} local_ctx(local_ctx const & other):object_ref(other) {} - local_ctx(local_ctx && other):object_ref(other) {} + local_ctx(local_ctx && other):object_ref(std::move(other)) {} local_ctx & operator=(local_ctx const & other) { object_ref::operator=(other); return *this; } - local_ctx & operator=(local_ctx && other) { object_ref::operator=(other); return *this; } + local_ctx & operator=(local_ctx && other) { object_ref::operator=(std::move(other)); return *this; } bool empty() const; diff --git a/stage0/src/kernel/replace_fn.cpp b/stage0/src/kernel/replace_fn.cpp index 0dde49600eb5..0a2b1030d12b 100644 --- a/stage0/src/kernel/replace_fn.cpp +++ b/stage0/src/kernel/replace_fn.cpp @@ -6,6 +6,7 @@ Author: Leonardo de Moura */ #include #include +#include #include #include "kernel/replace_fn.h" @@ -21,7 +22,7 @@ class replace_rec_fn { std::function(expr const &, unsigned)> m_f; bool m_use_cache; - expr save_result(expr const & e, unsigned offset, expr const & r, bool shared) { + expr save_result(expr const & e, unsigned offset, expr r, bool shared) { if (shared) m_cache.insert(mk_pair(mk_pair(e.raw(), offset), r)); return r; @@ -36,7 +37,7 @@ class replace_rec_fn { shared = true; } if (optional r = m_f(e, offset)) { - return save_result(e, offset, *r, shared); + return save_result(e, offset, std::move(*r), shared); } else { switch (e.kind()) { case expr_kind::Const: case expr_kind::Sort: diff --git a/stage0/src/library/compiler/specialize.cpp b/stage0/src/library/compiler/specialize.cpp index e6aa59c715f3..6874412c894e 100644 --- a/stage0/src/library/compiler/specialize.cpp +++ b/stage0/src/library/compiler/specialize.cpp @@ -92,10 +92,10 @@ class spec_info : public object_ref { object_ref(mk_cnstr(0, ns, ks)) {} spec_info():spec_info(names(), spec_arg_kinds()) {} spec_info(spec_info const & other):object_ref(other) {} - spec_info(spec_info && other):object_ref(other) {} + spec_info(spec_info && other):object_ref(std::move(other)) {} spec_info(b_obj_arg o, bool b):object_ref(o, b) {} spec_info & operator=(spec_info const & other) { object_ref::operator=(other); return *this; } - spec_info & operator=(spec_info && other) { object_ref::operator=(other); return *this; } + spec_info & operator=(spec_info && other) { object_ref::operator=(std::move(other)); return *this; } names const & get_mutual_decls() const { return static_cast(cnstr_get_ref(*this, 0)); } spec_arg_kinds const & get_arg_kinds() const { return static_cast(cnstr_get_ref(*this, 1)); } }; diff --git a/stage0/src/library/projection.h b/stage0/src/library/projection.h index 11e1abb2c83e..3ade6827e63c 100644 --- a/stage0/src/library/projection.h +++ b/stage0/src/library/projection.h @@ -24,13 +24,13 @@ class projection_info : public object_ref { projection_info(name const & c, unsigned nparams, unsigned i, bool inst_implicit); projection_info():projection_info(name(), 0, 0, false) {} projection_info(projection_info const & other):object_ref(other) {} - projection_info(projection_info && other):object_ref(other) {} + projection_info(projection_info && other):object_ref(std::move(other)) {} /* low-level constructors */ explicit projection_info(object * o):object_ref(o) {} explicit projection_info(b_obj_arg o, bool b):object_ref(o, b) {} explicit projection_info(object_ref const & o):object_ref(o) {} projection_info & operator=(projection_info const & other) { object_ref::operator=(other); return *this; } - projection_info & operator=(projection_info && other) { object_ref::operator=(other); return *this; } + projection_info & operator=(projection_info && other) { object_ref::operator=(std::move(other)); return *this; } name const & get_constructor() const { return static_cast(cnstr_get_ref(*this, 0)); } unsigned get_nparams() const { return static_cast(cnstr_get_ref(*this, 1)).get_small_value(); } unsigned get_i() const { return static_cast(cnstr_get_ref(*this, 2)).get_small_value(); } diff --git a/stage0/src/runtime/array_ref.h b/stage0/src/runtime/array_ref.h index a0f4db445042..468ca77340de 100644 --- a/stage0/src/runtime/array_ref.h +++ b/stage0/src/runtime/array_ref.h @@ -28,12 +28,12 @@ class array_ref : public object_ref { array_ref(b_obj_arg o, bool b):object_ref(o, b) {} array_ref():object_ref(array_mk_empty()) {} array_ref(array_ref const & other):object_ref(other) {} - array_ref(array_ref && other):object_ref(other) {} + array_ref(array_ref && other):object_ref(std::move(other)) {} array_ref(std::initializer_list const & elems):object_ref(to_array(elems)) {} array_ref(buffer const & elems):object_ref(to_array(elems)) {} array_ref & operator=(array_ref const & other) { object_ref::operator=(other); return *this; } - array_ref & operator=(array_ref && other) { object_ref::operator=(other); return *this; } + array_ref & operator=(array_ref && other) { object_ref::operator=(std::move(other)); return *this; } size_t size() const { return array_size(raw()); } diff --git a/stage0/src/runtime/list_ref.h b/stage0/src/runtime/list_ref.h index 19aa342a6f7c..1b5f163c566e 100644 --- a/stage0/src/runtime/list_ref.h +++ b/stage0/src/runtime/list_ref.h @@ -23,7 +23,7 @@ class list_ref : public object_ref { explicit list_ref(list_ref const * l) { if (l) *this = *l; } list_ref(T const & h, list_ref const & t):object_ref(mk_cnstr(1, h.raw(), t.raw())) { inc(h.raw()); inc(t.raw()); } list_ref(list_ref const & other):object_ref(other) {} - list_ref(list_ref && other):object_ref(other) {} + list_ref(list_ref && other):object_ref(std::move(other)) {} template list_ref(It const & begin, It const & end):list_ref() { auto it = end; while (it != begin) { @@ -35,7 +35,7 @@ class list_ref : public object_ref { list_ref(buffer const & b):list_ref(b.begin(), b.end()) {} list_ref & operator=(list_ref const & other) { object_ref::operator=(other); return *this; } - list_ref & operator=(list_ref && other) { object_ref::operator=(other); return *this; } + list_ref & operator=(list_ref && other) { object_ref::operator=(std::move(other)); return *this; } explicit operator bool() const { return !is_scalar(raw()); } friend bool is_nil(list_ref const & l) { return is_scalar(l.raw()); } diff --git a/stage0/src/runtime/option_ref.h b/stage0/src/runtime/option_ref.h index 187616e6c967..18c3f915ec4f 100644 --- a/stage0/src/runtime/option_ref.h +++ b/stage0/src/runtime/option_ref.h @@ -16,14 +16,14 @@ class option_ref : public object_ref { option_ref(b_obj_arg o, bool b):object_ref(o, b) {} option_ref():object_ref(box(0)) {} option_ref(option_ref const & other):object_ref(other) {} - option_ref(option_ref && other):object_ref(other) {} + option_ref(option_ref && other):object_ref(std::move(other)) {} explicit option_ref(T const & a):object_ref(mk_cnstr(1, a.raw())) { inc(a.raw()); } explicit option_ref(T const * a) { if (a) *this = option_ref(*a); } explicit option_ref(option_ref const * o) { if (o) *this = *o; } explicit option_ref(optional const & o):option_ref(o ? &*o : nullptr) {} option_ref & operator=(option_ref const & other) { object_ref::operator=(other); return *this; } - option_ref & operator=(option_ref && other) { object_ref::operator=(other); return *this; } + option_ref & operator=(option_ref && other) { object_ref::operator=(std::move(other)); return *this; } explicit operator bool() const { return !is_scalar(raw()); } optional get() const { return *this ? some(static_cast(cnstr_get_ref(*this, 0))) : optional(); } diff --git a/stage0/src/runtime/optional.h b/stage0/src/runtime/optional.h index cf3455fc548d..995fb3185938 100644 --- a/stage0/src/runtime/optional.h +++ b/stage0/src/runtime/optional.h @@ -31,13 +31,13 @@ class optional { } optional(optional && other):m_some(other.m_some) { if (other.m_some) - new (&m_value) T(std::forward(other.m_value)); + new (&m_value) T(std::move(other.m_value)); } explicit optional(T const & v):m_some(true) { new (&m_value) T(v); } explicit optional(T && v):m_some(true) { - new (&m_value) T(std::forward(v)); + new (&m_value) T(std::move(v)); } template explicit optional(Args&&... args):m_some(true) { @@ -84,7 +84,7 @@ class optional { m_value.~T(); m_some = other.m_some; if (m_some) - new (&m_value) T(std::forward(other.m_value)); + new (&m_value) T(std::move(other.m_value)); return *this; } optional& operator=(T const & other) { @@ -98,7 +98,7 @@ class optional { if (m_some) m_value.~T(); m_some = true; - new (&m_value) T(std::forward(other)); + new (&m_value) T(std::move(other)); return *this; } @@ -130,9 +130,9 @@ template<> class optional

{ \ public: \ optional():m_value(nullptr) {} \ optional(optional const & other):m_value(other.m_value) {} \ - optional(optional && other):m_value(std::forward

(other.m_value)) {} \ + optional(optional && other):m_value(std::move(other.m_value)) {} \ explicit optional(P const & v):m_value(v) {} \ - explicit optional(P && v):m_value(std::forward

(v)) {} \ + explicit optional(P && v):m_value(std::move(v)) {} \ \ explicit operator bool() const { return m_value.m_ptr != nullptr; } \ P const * operator->() const { lean_assert(m_value.m_ptr); return &m_value; } \ @@ -142,9 +142,9 @@ public: \ P const & value() const { lean_assert(m_value.m_ptr); return m_value; } \ P & value() { lean_assert(m_value.m_ptr); return m_value; } \ optional & operator=(optional const & other) { m_value = other.m_value; return *this; } \ - optional& operator=(optional && other) { m_value = std::forward

(other.m_value); return *this; } \ + optional& operator=(optional && other) { m_value = std::move(other.m_value); return *this; } \ optional& operator=(P const & other) { m_value = other; return *this; } \ - optional& operator=(P && other) { m_value = std::forward

(other); return *this; } \ + optional& operator=(P && other) { m_value = std::move(other); return *this; } \ friend bool operator==(optional const & o1, optional const & o2) { \ return static_cast(o1) == static_cast(o2) && (!o1 || o1.m_value == o2.m_value); \ } \ diff --git a/stage0/src/runtime/string_ref.h b/stage0/src/runtime/string_ref.h index a510ee837577..dd9fb726cb9a 100644 --- a/stage0/src/runtime/string_ref.h +++ b/stage0/src/runtime/string_ref.h @@ -17,11 +17,11 @@ class string_ref : public object_ref { explicit string_ref(std::string const & s):object_ref(mk_string(s)) {} explicit string_ref(sstream const & strm):string_ref(strm.str()) {} string_ref(string_ref const & other):object_ref(other) {} - string_ref(string_ref && other):object_ref(other) {} + string_ref(string_ref && other):object_ref(std::move(other)) {} explicit string_ref(obj_arg o):object_ref(o) {} string_ref(b_obj_arg o, bool b):object_ref(o, b) {} string_ref & operator=(string_ref const & other) { object_ref::operator=(other); return *this; } - string_ref & operator=(string_ref && other) { object_ref::operator=(other); return *this; } + string_ref & operator=(string_ref && other) { object_ref::operator=(std::move(other)); return *this; } /* Number of bytes used to store the string in UTF8. Remark: it does not include the null character added in the end to make it efficient to convert to C string. */ diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 658ab0874e68..0699845ba452 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); // toggling `parseQuotWithCurrentStage` may also require toggling the following option if macros/syntax // with custom precheck hooks were affected opts = opts.update({"quotPrecheck"}, true); diff --git a/stage0/src/util/kvmap.h b/stage0/src/util/kvmap.h index 172c4414726e..ff16d98328a8 100644 --- a/stage0/src/util/kvmap.h +++ b/stage0/src/util/kvmap.h @@ -30,9 +30,9 @@ class data_value : public object_ref { explicit data_value(name const & v):object_ref(mk_cnstr(static_cast(data_value_kind::Name), v.raw())) { inc(v.raw()); } data_value():data_value(false) {} data_value(data_value const & other):object_ref(other) {} - data_value(data_value && other):object_ref(other) {} + data_value(data_value && other):object_ref(std::move(other)) {} data_value & operator=(data_value const & other) { object_ref::operator=(other); return *this; } - data_value & operator=(data_value && other) { object_ref::operator=(other); return *this; } + data_value & operator=(data_value && other) { object_ref::operator=(std::move(other)); return *this; } data_value_kind kind() const { return static_cast(cnstr_tag(raw())); } string_ref const & get_string() const { lean_assert(kind() == data_value_kind::String); return static_cast(cnstr_get_ref(*this, 0)); } diff --git a/stage0/src/util/name.h b/stage0/src/util/name.h index 9c1d70e57525..f15fa1e29d0b 100644 --- a/stage0/src/util/name.h +++ b/stage0/src/util/name.h @@ -74,7 +74,7 @@ class name : public object_ref { name(std::string const & s):name(name(), string_ref(s)) {} name(string_ref const & s):name(name(), s) {} name(name const & other):object_ref(other) {} - name(name && other):object_ref(other) {} + name(name && other):object_ref(std::move(other)) {} /** \brief Create a hierarchical name using the given strings. Example: name{"foo", "bla", "tst"} creates the hierarchical @@ -97,7 +97,7 @@ class name : public object_ref { */ static name mk_internal_unique_name(); name & operator=(name const & other) { object_ref::operator=(other); return *this; } - name & operator=(name && other) { object_ref::operator=(other); return *this; } + name & operator=(name && other) { object_ref::operator=(std::move(other)); return *this; } static uint64_t hash(b_obj_arg n) { lean_assert(lean_name_hash(n) == lean_name_hash_exported_b(n)); return lean_name_hash(n); diff --git a/stage0/src/util/nat.h b/stage0/src/util/nat.h index e0bae8429a25..8035a9cd5f2b 100644 --- a/stage0/src/util/nat.h +++ b/stage0/src/util/nat.h @@ -27,10 +27,10 @@ class nat : public object_ref { static nat of_size_t(size_t v) { return nat(lean_usize_to_nat(v)); } nat(nat const & other):object_ref(other) {} - nat(nat && other):object_ref(other) {} + nat(nat && other):object_ref(std::move(other)) {} nat & operator=(nat const & other) { object_ref::operator=(other); return *this; } - nat & operator=(nat && other) { object_ref::operator=(other); return *this; } + nat & operator=(nat && other) { object_ref::operator=(std::move(other)); return *this; } bool is_small() const { return is_scalar(raw()); } size_t get_small_value() const { lean_assert(is_small()); return unbox(raw()); } bool is_zero() const { return is_small() && get_small_value() == 0; } diff --git a/stage0/src/util/shell.cpp b/stage0/src/util/shell.cpp index b5053f84ba5e..779c682b8976 100644 --- a/stage0/src/util/shell.cpp +++ b/stage0/src/util/shell.cpp @@ -310,7 +310,9 @@ options set_config_option(options const & opts, char const * in) { << "' cannot be set in the command line, use set_option command"); } } else { - throw lean::exception(lean::sstream() << "invalid -D parameter, unknown configuration option '" << opt << "'"); + // More options may be registered by imports, so we leave validating them to the Lean side. + // This (minor) duplication will be resolved when this file is rewritten in Lean. + return opts.update(opt, val.c_str()); } } diff --git a/stage0/stdlib/Init/Data/BitVec/Basic.c b/stage0/stdlib/Init/Data/BitVec/Basic.c index 99c16d7fa507..e8c0624f82d0 100644 --- a/stage0/stdlib/Init/Data/BitVec/Basic.c +++ b/stage0/stdlib/Init/Data/BitVec/Basic.c @@ -3498,11 +3498,9 @@ return x_10; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_nat_mul(x_1, x_4); -x_12 = l_BitVec_ofNat(x_11, x_4); -lean_dec(x_11); -return x_12; +lean_object* x_11; +x_11 = l_BitVec_nil___closed__1; +return x_11; } } } diff --git a/stage0/stdlib/Init/Data/BitVec/Lemmas.c b/stage0/stdlib/Init/Data/BitVec/Lemmas.c index 9a123a215cf9..5a395836b61b 100644 --- a/stage0/stdlib/Init/Data/BitVec/Lemmas.c +++ b/stage0/stdlib/Init/Data/BitVec/Lemmas.c @@ -13,12 +13,17 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_ofBoolListBE_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_BitVec_ofNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BitVec_intMax___boxed(lean_object*); LEAN_EXPORT lean_object* l_BitVec_intMax(lean_object*); lean_object* lean_nat_pow(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_ofBoolListBE_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_ofBoolListBE_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l_BitVec_intMax(lean_object* x_1) { @@ -83,6 +88,56 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_eq(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_1, x_7); +x_9 = lean_apply_2(x_4, x_8, x_2); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_apply_1(x_3, x_2); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg___boxed), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Init_Data_BitVec_Lemmas_0__BitVec_replicate_match__1_splitter(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} lean_object* initialize_Init_Data_Bool(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_BitVec_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Fin_Lemmas(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Init/Data/Int/Basic.c b/stage0/stdlib/Init/Data/Int/Basic.c index b540680c9f2e..90a5dc1e4ca0 100644 --- a/stage0/stdlib/Init/Data/Int/Basic.c +++ b/stage0/stdlib/Init/Data/Int/Basic.c @@ -52,6 +52,7 @@ lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_instCoeHTCTIntOfIntCast___rarg(lean_object*); LEAN_EXPORT lean_object* l_Int_pow(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int___aux__Init__Data__Int__Basic______unexpand__Int__negSucc__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Int_instNatPow; LEAN_EXPORT lean_object* l_Int_decLe___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_decLt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_negSucc___boxed(lean_object*); @@ -64,6 +65,7 @@ LEAN_EXPORT lean_object* l_Int_cast___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_instLTInt; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Int_instNatPow___closed__1; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Int___aux__Init__Data__Int__Basic______unexpand__Int__negSucc__1___closed__1; @@ -78,7 +80,6 @@ LEAN_EXPORT lean_object* l_Int_subNatNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_ofNat___boxed(lean_object*); static lean_object* l_Int___aux__Init__Data__Int__Basic______macroRules__Int__term_x2d_x5b___x2b1_x5d__1___closed__3; LEAN_EXPORT lean_object* l_Int_subNatNat___boxed(lean_object*, lean_object*); -static lean_object* l_Int_instHPowNat___closed__1; lean_object* lean_nat_abs(lean_object*); lean_object* lean_int_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_instDecidableEq___boxed(lean_object*, lean_object*); @@ -94,7 +95,6 @@ LEAN_EXPORT lean_object* l_Int_neg___boxed(lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_instSub; static lean_object* l_Int_term_x2d_x5b___x2b1_x5d___closed__4; -LEAN_EXPORT lean_object* l_Int_instHPowNat; LEAN_EXPORT lean_object* l_Int_mul___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_toNat(lean_object*); uint8_t lean_int_dec_lt(lean_object*, lean_object*); @@ -1075,7 +1075,7 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l_Int_instHPowNat___closed__1() { +static lean_object* _init_l_Int_instNatPow___closed__1() { _start: { lean_object* x_1; @@ -1083,11 +1083,11 @@ x_1 = lean_alloc_closure((void*)(l_Int_pow___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Int_instHPowNat() { +static lean_object* _init_l_Int_instNatPow() { _start: { lean_object* x_1; -x_1 = l_Int_instHPowNat___closed__1; +x_1 = l_Int_instNatPow___closed__1; return x_1; } } @@ -1324,10 +1324,10 @@ l_Int_sign___closed__2 = _init_l_Int_sign___closed__2(); lean_mark_persistent(l_Int_sign___closed__2); l_Int_instDvd = _init_l_Int_instDvd(); lean_mark_persistent(l_Int_instDvd); -l_Int_instHPowNat___closed__1 = _init_l_Int_instHPowNat___closed__1(); -lean_mark_persistent(l_Int_instHPowNat___closed__1); -l_Int_instHPowNat = _init_l_Int_instHPowNat(); -lean_mark_persistent(l_Int_instHPowNat); +l_Int_instNatPow___closed__1 = _init_l_Int_instNatPow___closed__1(); +lean_mark_persistent(l_Int_instNatPow___closed__1); +l_Int_instNatPow = _init_l_Int_instNatPow(); +lean_mark_persistent(l_Int_instNatPow); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lake/Build/Module.c b/stage0/stdlib/Lake/Build/Module.c index 532f294f28a0..30c2d4375874 100644 --- a/stage0/stdlib/Lake/Build/Module.c +++ b/stage0/stdlib/Lake/Build/Module.c @@ -198,7 +198,6 @@ lean_object* l_Lake_withRegisterJob___rarg(lean_object*, lean_object*, uint8_t, static lean_object* l_Lake_initModuleFacetConfigs___closed__1; static lean_object* l_Lake_Module_oExportFacetConfig___closed__2; LEAN_EXPORT lean_object* l_Lake_Module_bcFacetConfig; -lean_object* l_Lean_moduleExists_go(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_initModuleFacetConfigs___closed__16; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lake_Module_recComputePrecompileImports___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lake_Module_precompileImportsFacetConfig___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -245,7 +244,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDeps___ lean_object* l_IO_FS_Stream_ofBuffer(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lake_Module_recBuildDeps___spec__10(lean_object*, lean_object*); static lean_object* l_Lake_Module_coExportFacetConfig___closed__4; -LEAN_EXPORT lean_object* l_Lake_Module_recParseImports___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_initModuleFacetConfigs___closed__17; LEAN_EXPORT lean_object* l_Lake_Module_recBuildLeanCToONoExport(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lake_buildFileUnlessUpToDate___spec__5(lean_object*, lean_object*); @@ -277,7 +275,6 @@ LEAN_EXPORT lean_object* l_Lake_Module_bcoFacetConfig; LEAN_EXPORT lean_object* l_Lake_OrdHashSet_appendArray___at_Lake_collectImportsAux___spec__1(lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lake_Module_recComputeTransImports(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lake_Module_recParseImports___closed__2; static lean_object* l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__3; LEAN_EXPORT lean_object* l_Lake_Module_recBuildDeps___lambda__1___closed__3___boxed__const__1; static lean_object* l_Lake_Module_coExportFacetConfig___closed__1; @@ -392,7 +389,6 @@ LEAN_EXPORT lean_object* l_Lake_Module_recBuildLean___lambda__5(lean_object*, le lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lake_Module_recBuildDeps___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lake_Module_oExportFacetConfig; -LEAN_EXPORT lean_object* l_Lake_Module_recParseImports___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lake_Module_recBuildDynlib___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_Module_coNoExportFacetConfig___closed__3; lean_object* l_Lean_modToFilePath(lean_object*, lean_object*, lean_object*); @@ -415,7 +411,6 @@ lean_object* l_Lake_fetchFileTrace___boxed(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lake_Module_oNoExportFacetConfig; static lean_object* l_Lake_Module_depsFacetConfig___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDeps___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lake_Module_recParseImports___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lake_Module_recComputePrecompileImports(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_OrdHashSet_insert___at_Lake_Module_recBuildDeps___spec__6___closed__1; static lean_object* l_Lake_Module_dynlibFacetConfig___closed__1; @@ -431,6 +426,7 @@ LEAN_EXPORT lean_object* l_Functor_discard___at_Lake_Module_dynlibFacetConfig___ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___lambda__3(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_replace___at_Lake_Module_recParseImports___spec__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; lean_object* l_Lake_BuildType_leancArgs(uint8_t); static lean_object* l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDeps___spec__1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lake_Module_recBuildDeps___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*); @@ -2881,7 +2877,7 @@ return x_59; } } } -static lean_object* _init_l_Lake_Module_recParseImports___lambda__1___closed__1() { +static lean_object* _init_l_Lake_Module_recParseImports___closed__1() { _start: { lean_object* x_1; @@ -2889,465 +2885,465 @@ x_1 = lean_mk_string_unchecked("lean", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lake_Module_recParseImports___lambda__1(lean_object* x_1, lean_object* x_2, lean_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_Lake_Module_recParseImports(lean_object* x_1, lean_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_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_271; -x_193 = lean_ctor_get(x_1, 0); +lean_object* x_8; lean_object* x_9; 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_270; +x_192 = lean_ctor_get(x_1, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_192, 0); lean_inc(x_193); x_194 = lean_ctor_get(x_193, 0); lean_inc(x_194); -x_195 = lean_ctor_get(x_194, 0); +x_195 = lean_ctor_get(x_193, 2); lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 2); +lean_dec(x_193); +x_196 = lean_ctor_get(x_195, 7); lean_inc(x_196); -lean_dec(x_194); -x_197 = lean_ctor_get(x_196, 7); -lean_inc(x_197); -lean_dec(x_196); -x_198 = l_System_FilePath_join(x_195, x_197); -x_199 = lean_ctor_get(x_193, 1); +lean_dec(x_195); +x_197 = l_System_FilePath_join(x_194, x_196); +x_198 = lean_ctor_get(x_192, 1); +lean_inc(x_198); +lean_dec(x_192); +x_199 = lean_ctor_get(x_198, 2); lean_inc(x_199); -lean_dec(x_193); -x_200 = lean_ctor_get(x_199, 2); -lean_inc(x_200); -lean_dec(x_199); -x_201 = l_System_FilePath_join(x_198, x_200); -x_202 = lean_ctor_get(x_1, 1); -lean_inc(x_202); +lean_dec(x_198); +x_200 = l_System_FilePath_join(x_197, x_199); +x_201 = lean_ctor_get(x_1, 1); +lean_inc(x_201); lean_dec(x_1); -x_203 = l_Lake_Module_recParseImports___lambda__1___closed__1; -x_204 = l_Lean_modToFilePath(x_201, x_202, x_203); -lean_dec(x_201); -x_271 = l_IO_FS_readFile(x_204, x_8); -if (lean_obj_tag(x_271) == 0) -{ -uint8_t x_272; -x_272 = !lean_is_exclusive(x_271); -if (x_272 == 0) +x_202 = l_Lake_Module_recParseImports___closed__1; +x_203 = l_Lean_modToFilePath(x_200, x_201, x_202); +lean_dec(x_200); +x_270 = l_IO_FS_readFile(x_203, x_7); +if (lean_obj_tag(x_270) == 0) +{ +uint8_t x_271; +x_271 = !lean_is_exclusive(x_270); +if (x_271 == 0) { -lean_object* x_273; lean_object* x_274; -x_273 = lean_ctor_get(x_271, 1); -lean_ctor_set(x_271, 1, x_6); -x_274 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_274, 0, x_271); -lean_ctor_set(x_274, 1, x_7); -x_205 = x_274; -x_206 = x_273; -goto block_270; +lean_object* x_272; lean_object* x_273; +x_272 = lean_ctor_get(x_270, 1); +lean_ctor_set(x_270, 1, x_5); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_6); +x_204 = x_273; +x_205 = x_272; +goto block_269; } else { -lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_275 = lean_ctor_get(x_271, 0); -x_276 = lean_ctor_get(x_271, 1); -lean_inc(x_276); +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_274 = lean_ctor_get(x_270, 0); +x_275 = lean_ctor_get(x_270, 1); lean_inc(x_275); -lean_dec(x_271); +lean_inc(x_274); +lean_dec(x_270); +x_276 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_276, 0, x_274); +lean_ctor_set(x_276, 1, x_5); x_277 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 0, x_276); lean_ctor_set(x_277, 1, x_6); -x_278 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_278, 0, x_277); -lean_ctor_set(x_278, 1, x_7); -x_205 = x_278; -x_206 = x_276; -goto block_270; +x_204 = x_277; +x_205 = x_275; +goto block_269; } } else { -uint8_t x_279; -x_279 = !lean_is_exclusive(x_271); -if (x_279 == 0) -{ -lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; -x_280 = lean_ctor_get(x_271, 0); -x_281 = lean_ctor_get(x_271, 1); -x_282 = lean_io_error_to_string(x_280); -x_283 = 3; -x_284 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set_uint8(x_284, sizeof(void*)*1, x_283); -x_285 = lean_array_get_size(x_6); -x_286 = lean_array_push(x_6, x_284); -lean_ctor_set(x_271, 1, x_286); -lean_ctor_set(x_271, 0, x_285); -x_287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_287, 0, x_271); -lean_ctor_set(x_287, 1, x_7); -x_205 = x_287; -x_206 = x_281; -goto block_270; -} -else +uint8_t x_278; +x_278 = !lean_is_exclusive(x_270); +if (x_278 == 0) { -lean_object* x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; -x_288 = lean_ctor_get(x_271, 0); -x_289 = lean_ctor_get(x_271, 1); -lean_inc(x_289); +lean_object* x_279; lean_object* x_280; lean_object* x_281; uint8_t x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_279 = lean_ctor_get(x_270, 0); +x_280 = lean_ctor_get(x_270, 1); +x_281 = lean_io_error_to_string(x_279); +x_282 = 3; +x_283 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_283, 0, x_281); +lean_ctor_set_uint8(x_283, sizeof(void*)*1, x_282); +x_284 = lean_array_get_size(x_5); +x_285 = lean_array_push(x_5, x_283); +lean_ctor_set(x_270, 1, x_285); +lean_ctor_set(x_270, 0, x_284); +x_286 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_286, 0, x_270); +lean_ctor_set(x_286, 1, x_6); +x_204 = x_286; +x_205 = x_280; +goto block_269; +} +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; +x_287 = lean_ctor_get(x_270, 0); +x_288 = lean_ctor_get(x_270, 1); lean_inc(x_288); -lean_dec(x_271); -x_290 = lean_io_error_to_string(x_288); -x_291 = 3; -x_292 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_292, 0, x_290); -lean_ctor_set_uint8(x_292, sizeof(void*)*1, x_291); -x_293 = lean_array_get_size(x_6); -x_294 = lean_array_push(x_6, x_292); -x_295 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_295, 0, x_293); -lean_ctor_set(x_295, 1, x_294); -x_296 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_296, 0, x_295); -lean_ctor_set(x_296, 1, x_7); -x_205 = x_296; -x_206 = x_289; -goto block_270; +lean_inc(x_287); +lean_dec(x_270); +x_289 = lean_io_error_to_string(x_287); +x_290 = 3; +x_291 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set_uint8(x_291, sizeof(void*)*1, x_290); +x_292 = lean_array_get_size(x_5); +x_293 = lean_array_push(x_5, x_291); +x_294 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); +x_295 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_295, 0, x_294); +lean_ctor_set(x_295, 1, x_6); +x_204 = x_295; +x_205 = x_288; +goto block_269; } } -block_192: +block_191: { -lean_object* x_11; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) +lean_object* x_10; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_9, 1); -x_14 = lean_ctor_get(x_9, 0); -lean_dec(x_14); -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_8, 1); +x_13 = lean_ctor_get(x_8, 0); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -x_18 = lean_array_get_size(x_16); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +x_17 = lean_array_get_size(x_15); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_lt(x_18, x_17); +if (x_19 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_5); +lean_object* x_20; lean_object* x_21; +lean_dec(x_17); +lean_dec(x_15); lean_dec(x_4); lean_dec(x_3); -x_21 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; -lean_ctor_set(x_11, 0, x_21); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_9); -lean_ctor_set(x_22, 1, x_10); -return x_22; +lean_dec(x_2); +x_20 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +lean_ctor_set(x_10, 0, x_20); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_9); +return x_21; } else { -uint8_t x_23; -x_23 = lean_nat_dec_le(x_18, x_18); -if (x_23 == 0) +uint8_t x_22; +x_22 = lean_nat_dec_le(x_17, x_17); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_5); +lean_object* x_23; lean_object* x_24; +lean_dec(x_17); +lean_dec(x_15); lean_dec(x_4); lean_dec(x_3); -x_24 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; -lean_ctor_set(x_11, 0, x_24); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_9); -lean_ctor_set(x_25, 1, x_10); -return x_25; +lean_dec(x_2); +x_23 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +lean_ctor_set(x_10, 0, x_23); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_8); +lean_ctor_set(x_24, 1, x_9); +return x_24; } else { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; -lean_free_object(x_11); -lean_free_object(x_9); -x_26 = 0; -x_27 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_28 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; -x_29 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_16, x_26, x_27, x_28, x_3, x_4, x_5, x_17, x_13, x_10); -lean_dec(x_16); -if (lean_obj_tag(x_29) == 0) +size_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_10); +lean_free_object(x_8); +x_25 = 0; +x_26 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_27 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; +x_28 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_15, x_25, x_26, x_27, x_2, x_3, x_4, x_16, x_12, x_9); +lean_dec(x_15); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_30; lean_object* x_31; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) +if (lean_obj_tag(x_30) == 0) { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_29); -if (x_32 == 0) +uint8_t x_31; +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_29, 0); -lean_dec(x_33); -x_34 = !lean_is_exclusive(x_30); -if (x_34 == 0) +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_28, 0); +lean_dec(x_32); +x_33 = !lean_is_exclusive(x_29); +if (x_33 == 0) { -lean_object* x_35; uint8_t x_36; -x_35 = lean_ctor_get(x_30, 0); -lean_dec(x_35); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_29, 0); +lean_dec(x_34); +x_35 = !lean_is_exclusive(x_30); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_31, 0); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -lean_ctor_set(x_31, 0, x_38); -return x_29; +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_30, 0); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +lean_ctor_set(x_30, 0, x_37); +return x_28; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_31, 0); -x_40 = lean_ctor_get(x_31, 1); -lean_inc(x_40); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); lean_inc(x_39); -lean_dec(x_31); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -lean_ctor_set(x_30, 0, x_42); -return x_29; +lean_inc(x_38); +lean_dec(x_30); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_29, 0, x_41); +return x_28; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_43 = lean_ctor_get(x_30, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_42 = lean_ctor_get(x_29, 1); +lean_inc(x_42); +lean_dec(x_29); +x_43 = lean_ctor_get(x_30, 0); lean_inc(x_43); -lean_dec(x_30); -x_44 = lean_ctor_get(x_31, 0); +x_44 = lean_ctor_get(x_30, 1); lean_inc(x_44); -x_45 = lean_ctor_get(x_31, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_46 = x_31; +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_45 = x_30; } else { - lean_dec_ref(x_31); - x_46 = lean_box(0); + lean_dec_ref(x_30); + x_45 = lean_box(0); } -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -if (lean_is_scalar(x_46)) { - x_48 = lean_alloc_ctor(0, 2, 0); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +if (lean_is_scalar(x_45)) { + x_47 = lean_alloc_ctor(0, 2, 0); } else { - x_48 = x_46; + x_47 = x_45; } +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_44); +x_48 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_45); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_43); -lean_ctor_set(x_29, 0, x_49); -return x_29; +lean_ctor_set(x_48, 1, x_42); +lean_ctor_set(x_28, 0, x_48); +return x_28; } } 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_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_49 = lean_ctor_get(x_28, 1); +lean_inc(x_49); +lean_dec(x_28); x_50 = lean_ctor_get(x_29, 1); lean_inc(x_50); -lean_dec(x_29); -x_51 = lean_ctor_get(x_30, 1); -lean_inc(x_51); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_51 = x_29; +} else { + lean_dec_ref(x_29); + x_51 = lean_box(0); +} +x_52 = lean_ctor_get(x_30, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_30, 1); +lean_inc(x_53); if (lean_is_exclusive(x_30)) { lean_ctor_release(x_30, 0); lean_ctor_release(x_30, 1); - x_52 = x_30; + x_54 = x_30; } else { lean_dec_ref(x_30); - x_52 = lean_box(0); + x_54 = lean_box(0); } -x_53 = lean_ctor_get(x_31, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_31, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_55 = x_31; +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +lean_dec(x_52); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_31); - x_55 = lean_box(0); + x_56 = x_54; } -x_56 = lean_ctor_get(x_53, 1); -lean_inc(x_56); -lean_dec(x_53); -if (lean_is_scalar(x_55)) { +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +if (lean_is_scalar(x_51)) { x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_55; + x_57 = x_51; } lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -if (lean_is_scalar(x_52)) { - x_58 = lean_alloc_ctor(0, 2, 0); -} else { - x_58 = x_52; -} +lean_ctor_set(x_57, 1, x_50); +x_58 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_51); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_50); -return x_59; +lean_ctor_set(x_58, 1, x_49); +return x_58; } } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_29); -if (x_60 == 0) +uint8_t x_59; +x_59 = !lean_is_exclusive(x_28); +if (x_59 == 0) { -lean_object* x_61; uint8_t x_62; -x_61 = lean_ctor_get(x_29, 0); -lean_dec(x_61); -x_62 = !lean_is_exclusive(x_30); -if (x_62 == 0) +lean_object* x_60; uint8_t x_61; +x_60 = lean_ctor_get(x_28, 0); +lean_dec(x_60); +x_61 = !lean_is_exclusive(x_29); +if (x_61 == 0) { -lean_object* x_63; uint8_t x_64; -x_63 = lean_ctor_get(x_30, 0); -lean_dec(x_63); -x_64 = !lean_is_exclusive(x_31); -if (x_64 == 0) +lean_object* x_62; uint8_t x_63; +x_62 = lean_ctor_get(x_29, 0); +lean_dec(x_62); +x_63 = !lean_is_exclusive(x_30); +if (x_63 == 0) { -return x_29; +return x_28; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_31, 0); -x_66 = lean_ctor_get(x_31, 1); -lean_inc(x_66); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_30, 0); +x_65 = lean_ctor_get(x_30, 1); lean_inc(x_65); -lean_dec(x_31); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_30, 0, x_67); -return x_29; +lean_inc(x_64); +lean_dec(x_30); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_29, 0, x_66); +return x_28; } } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_68 = lean_ctor_get(x_30, 1); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_67 = lean_ctor_get(x_29, 1); +lean_inc(x_67); +lean_dec(x_29); +x_68 = lean_ctor_get(x_30, 0); lean_inc(x_68); -lean_dec(x_30); -x_69 = lean_ctor_get(x_31, 0); +x_69 = lean_ctor_get(x_30, 1); lean_inc(x_69); -x_70 = lean_ctor_get(x_31, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_71 = x_31; +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_70 = x_30; } else { - lean_dec_ref(x_31); - x_71 = lean_box(0); + lean_dec_ref(x_30); + x_70 = lean_box(0); } -if (lean_is_scalar(x_71)) { - x_72 = lean_alloc_ctor(1, 2, 0); -} else { - x_72 = x_71; +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(1, 2, 0); +} else { + x_71 = x_70; } -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_70); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_68); -lean_ctor_set(x_29, 0, x_73); -return x_29; +lean_ctor_set(x_71, 0, x_68); +lean_ctor_set(x_71, 1, x_69); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_67); +lean_ctor_set(x_28, 0, x_72); +return x_28; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_73 = lean_ctor_get(x_28, 1); +lean_inc(x_73); +lean_dec(x_28); x_74 = lean_ctor_get(x_29, 1); lean_inc(x_74); -lean_dec(x_29); -x_75 = lean_ctor_get(x_30, 1); -lean_inc(x_75); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_75 = x_29; +} else { + lean_dec_ref(x_29); + x_75 = lean_box(0); +} +x_76 = lean_ctor_get(x_30, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_30, 1); +lean_inc(x_77); if (lean_is_exclusive(x_30)) { lean_ctor_release(x_30, 0); lean_ctor_release(x_30, 1); - x_76 = x_30; + x_78 = x_30; } else { lean_dec_ref(x_30); - x_76 = lean_box(0); -} -x_77 = lean_ctor_get(x_31, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_31, 1); -lean_inc(x_78); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_79 = x_31; -} else { - lean_dec_ref(x_31); - x_79 = lean_box(0); + x_78 = lean_box(0); } -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_80 = x_79; + x_79 = x_78; } -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_78); -if (lean_is_scalar(x_76)) { - x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +if (lean_is_scalar(x_75)) { + x_80 = lean_alloc_ctor(0, 2, 0); } else { - x_81 = x_76; + x_80 = x_75; } +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_74); +x_81 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_75); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_74); -return x_82; +lean_ctor_set(x_81, 1, x_73); +return x_81; } } } else { -uint8_t x_83; -x_83 = !lean_is_exclusive(x_29); -if (x_83 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_28); +if (x_82 == 0) { -return x_29; +return x_28; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_29, 0); -x_85 = lean_ctor_get(x_29, 1); -lean_inc(x_85); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_28, 0); +x_84 = lean_ctor_get(x_28, 1); lean_inc(x_84); -lean_dec(x_29); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -return x_86; +lean_inc(x_83); +lean_dec(x_28); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; } } } @@ -3355,215 +3351,215 @@ return x_86; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_87 = lean_ctor_get(x_11, 0); -x_88 = lean_ctor_get(x_11, 1); -lean_inc(x_88); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_86 = lean_ctor_get(x_10, 0); +x_87 = lean_ctor_get(x_10, 1); lean_inc(x_87); -lean_dec(x_11); -x_89 = lean_array_get_size(x_87); -x_90 = lean_unsigned_to_nat(0u); -x_91 = lean_nat_dec_lt(x_90, x_89); -if (x_91 == 0) +lean_inc(x_86); +lean_dec(x_10); +x_88 = lean_array_get_size(x_86); +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_nat_dec_lt(x_89, x_88); +if (x_90 == 0) { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_89); -lean_dec(x_87); -lean_dec(x_5); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_88); +lean_dec(x_86); lean_dec(x_4); lean_dec(x_3); -x_92 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +lean_dec(x_2); +x_91 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_87); +lean_ctor_set(x_8, 0, x_92); x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_88); -lean_ctor_set(x_9, 0, x_93); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_9); -lean_ctor_set(x_94, 1, x_10); -return x_94; +lean_ctor_set(x_93, 0, x_8); +lean_ctor_set(x_93, 1, x_9); +return x_93; } else { -uint8_t x_95; -x_95 = lean_nat_dec_le(x_89, x_89); -if (x_95 == 0) +uint8_t x_94; +x_94 = lean_nat_dec_le(x_88, x_88); +if (x_94 == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_89); -lean_dec(x_87); -lean_dec(x_5); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_88); +lean_dec(x_86); lean_dec(x_4); lean_dec(x_3); -x_96 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +lean_dec(x_2); +x_95 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_87); +lean_ctor_set(x_8, 0, x_96); x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_88); -lean_ctor_set(x_9, 0, x_97); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_9); -lean_ctor_set(x_98, 1, x_10); -return x_98; +lean_ctor_set(x_97, 0, x_8); +lean_ctor_set(x_97, 1, x_9); +return x_97; } else { -size_t x_99; size_t x_100; lean_object* x_101; lean_object* x_102; -lean_free_object(x_9); -x_99 = 0; -x_100 = lean_usize_of_nat(x_89); -lean_dec(x_89); -x_101 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; -x_102 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_87, x_99, x_100, x_101, x_3, x_4, x_5, x_88, x_13, x_10); -lean_dec(x_87); -if (lean_obj_tag(x_102) == 0) +size_t x_98; size_t x_99; lean_object* x_100; lean_object* x_101; +lean_free_object(x_8); +x_98 = 0; +x_99 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_100 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; +x_101 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_86, x_98, x_99, x_100, x_2, x_3, x_4, x_87, x_12, x_9); +lean_dec(x_86); +if (lean_obj_tag(x_101) == 0) { -lean_object* x_103; lean_object* x_104; +lean_object* x_102; lean_object* x_103; +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -if (lean_obj_tag(x_104) == 0) +if (lean_obj_tag(x_103) == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_104 = lean_ctor_get(x_101, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_105 = x_101; +} else { + lean_dec_ref(x_101); + x_105 = lean_box(0); +} +x_106 = lean_ctor_get(x_102, 1); +lean_inc(x_106); if (lean_is_exclusive(x_102)) { lean_ctor_release(x_102, 0); lean_ctor_release(x_102, 1); - x_106 = x_102; + x_107 = x_102; } else { lean_dec_ref(x_102); - x_106 = lean_box(0); + x_107 = lean_box(0); } -x_107 = lean_ctor_get(x_103, 1); -lean_inc(x_107); +x_108 = lean_ctor_get(x_103, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_103, 1); +lean_inc(x_109); if (lean_is_exclusive(x_103)) { lean_ctor_release(x_103, 0); lean_ctor_release(x_103, 1); - x_108 = x_103; + x_110 = x_103; } else { lean_dec_ref(x_103); - x_108 = lean_box(0); + x_110 = lean_box(0); } -x_109 = lean_ctor_get(x_104, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_104, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_111 = x_104; +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +lean_dec(x_108); +if (lean_is_scalar(x_110)) { + x_112 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_104); - x_111 = lean_box(0); + x_112 = x_110; } -x_112 = lean_ctor_get(x_109, 1); -lean_inc(x_112); -lean_dec(x_109); -if (lean_is_scalar(x_111)) { +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +if (lean_is_scalar(x_107)) { x_113 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_111; + x_113 = x_107; } lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_110); -if (lean_is_scalar(x_108)) { +lean_ctor_set(x_113, 1, x_106); +if (lean_is_scalar(x_105)) { x_114 = lean_alloc_ctor(0, 2, 0); } else { - x_114 = x_108; + x_114 = x_105; } lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_107); -if (lean_is_scalar(x_106)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_106; -} -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_105); -return x_115; +lean_ctor_set(x_114, 1, x_104); +return x_114; } 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; -x_116 = lean_ctor_get(x_102, 1); -lean_inc(x_116); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_115 = lean_ctor_get(x_101, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_116 = x_101; +} else { + lean_dec_ref(x_101); + x_116 = lean_box(0); +} +x_117 = lean_ctor_get(x_102, 1); +lean_inc(x_117); if (lean_is_exclusive(x_102)) { lean_ctor_release(x_102, 0); lean_ctor_release(x_102, 1); - x_117 = x_102; + x_118 = x_102; } else { lean_dec_ref(x_102); - x_117 = lean_box(0); + x_118 = lean_box(0); } -x_118 = lean_ctor_get(x_103, 1); -lean_inc(x_118); +x_119 = lean_ctor_get(x_103, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_103, 1); +lean_inc(x_120); if (lean_is_exclusive(x_103)) { lean_ctor_release(x_103, 0); lean_ctor_release(x_103, 1); - x_119 = x_103; + x_121 = x_103; } else { lean_dec_ref(x_103); - x_119 = lean_box(0); + x_121 = lean_box(0); } -x_120 = lean_ctor_get(x_104, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_104, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_122 = x_104; +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_104); - x_122 = lean_box(0); + x_122 = x_121; } -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +if (lean_is_scalar(x_118)) { + x_123 = lean_alloc_ctor(0, 2, 0); } else { - x_123 = x_122; + x_123 = x_118; } -lean_ctor_set(x_123, 0, x_120); -lean_ctor_set(x_123, 1, x_121); -if (lean_is_scalar(x_119)) { +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_117); +if (lean_is_scalar(x_116)) { x_124 = lean_alloc_ctor(0, 2, 0); } else { - x_124 = x_119; + x_124 = x_116; } lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_118); -if (lean_is_scalar(x_117)) { - x_125 = lean_alloc_ctor(0, 2, 0); -} else { - x_125 = x_117; -} -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_116); -return x_125; +lean_ctor_set(x_124, 1, x_115); +return x_124; } } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_126 = lean_ctor_get(x_102, 0); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_125 = lean_ctor_get(x_101, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_101, 1); lean_inc(x_126); -x_127 = lean_ctor_get(x_102, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_128 = x_102; +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_127 = x_101; } else { - lean_dec_ref(x_102); - x_128 = lean_box(0); + lean_dec_ref(x_101); + x_127 = lean_box(0); } -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(1, 2, 0); } else { - x_129 = x_128; + x_128 = x_127; } -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_127); -return x_129; +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +return x_128; } } } @@ -3571,995 +3567,574 @@ return x_129; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_130 = lean_ctor_get(x_9, 1); +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; +x_129 = lean_ctor_get(x_8, 1); +lean_inc(x_129); +lean_dec(x_8); +x_130 = lean_ctor_get(x_10, 0); lean_inc(x_130); -lean_dec(x_9); -x_131 = lean_ctor_get(x_11, 0); +x_131 = lean_ctor_get(x_10, 1); lean_inc(x_131); -x_132 = lean_ctor_get(x_11, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - lean_ctor_release(x_11, 1); - x_133 = x_11; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_132 = x_10; } else { - lean_dec_ref(x_11); - x_133 = lean_box(0); + lean_dec_ref(x_10); + x_132 = lean_box(0); } -x_134 = lean_array_get_size(x_131); -x_135 = lean_unsigned_to_nat(0u); -x_136 = lean_nat_dec_lt(x_135, x_134); -if (x_136 == 0) +x_133 = lean_array_get_size(x_130); +x_134 = lean_unsigned_to_nat(0u); +x_135 = lean_nat_dec_lt(x_134, x_133); +if (x_135 == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_134); -lean_dec(x_131); -lean_dec(x_5); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_133); +lean_dec(x_130); lean_dec(x_4); lean_dec(x_3); -x_137 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; -if (lean_is_scalar(x_133)) { - x_138 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_2); +x_136 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +if (lean_is_scalar(x_132)) { + x_137 = lean_alloc_ctor(0, 2, 0); } else { - x_138 = x_133; + x_137 = x_132; } +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_131); +x_138 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_132); +lean_ctor_set(x_138, 1, x_129); x_139 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_130); -x_140 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_10); -return x_140; +lean_ctor_set(x_139, 1, x_9); +return x_139; } else { -uint8_t x_141; -x_141 = lean_nat_dec_le(x_134, x_134); -if (x_141 == 0) +uint8_t x_140; +x_140 = lean_nat_dec_le(x_133, x_133); +if (x_140 == 0) { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -lean_dec(x_134); -lean_dec(x_131); -lean_dec(x_5); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_133); +lean_dec(x_130); lean_dec(x_4); lean_dec(x_3); -x_142 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; -if (lean_is_scalar(x_133)) { - x_143 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_2); +x_141 = l_Array_forInUnsafe_loop___at_Lake_recBuildExternDynlibs___spec__3___closed__1; +if (lean_is_scalar(x_132)) { + x_142 = lean_alloc_ctor(0, 2, 0); } else { - x_143 = x_133; + x_142 = x_132; } +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_131); +x_143 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_132); +lean_ctor_set(x_143, 1, x_129); x_144 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_130); -x_145 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_10); -return x_145; +lean_ctor_set(x_144, 1, x_9); +return x_144; } else { -size_t x_146; size_t x_147; lean_object* x_148; lean_object* x_149; +size_t x_145; size_t x_146; lean_object* x_147; lean_object* x_148; +lean_dec(x_132); +x_145 = 0; +x_146 = lean_usize_of_nat(x_133); lean_dec(x_133); -x_146 = 0; -x_147 = lean_usize_of_nat(x_134); -lean_dec(x_134); -x_148 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; -x_149 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_131, x_146, x_147, x_148, x_3, x_4, x_5, x_132, x_130, x_10); -lean_dec(x_131); -if (lean_obj_tag(x_149) == 0) +x_147 = l_Lake_OrdHashSet_empty___at_Lake_OrdModuleSet_empty___spec__1; +x_148 = l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12(x_130, x_145, x_146, x_147, x_2, x_3, x_4, x_131, x_129, x_9); +lean_dec(x_130); +if (lean_obj_tag(x_148) == 0) { -lean_object* x_150; lean_object* x_151; +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); x_150 = lean_ctor_get(x_149, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -if (lean_obj_tag(x_151) == 0) +if (lean_obj_tag(x_150) == 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; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_152 = lean_ctor_get(x_149, 1); -lean_inc(x_152); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_151 = lean_ctor_get(x_148, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_152 = x_148; +} else { + lean_dec_ref(x_148); + x_152 = lean_box(0); +} +x_153 = lean_ctor_get(x_149, 1); +lean_inc(x_153); if (lean_is_exclusive(x_149)) { lean_ctor_release(x_149, 0); lean_ctor_release(x_149, 1); - x_153 = x_149; + x_154 = x_149; } else { lean_dec_ref(x_149); - x_153 = lean_box(0); + x_154 = lean_box(0); } -x_154 = lean_ctor_get(x_150, 1); -lean_inc(x_154); +x_155 = lean_ctor_get(x_150, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_150, 1); +lean_inc(x_156); if (lean_is_exclusive(x_150)) { lean_ctor_release(x_150, 0); lean_ctor_release(x_150, 1); - x_155 = x_150; + x_157 = x_150; } else { lean_dec_ref(x_150); - x_155 = lean_box(0); + x_157 = lean_box(0); } -x_156 = lean_ctor_get(x_151, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_151, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_158 = x_151; +x_158 = lean_ctor_get(x_155, 1); +lean_inc(x_158); +lean_dec(x_155); +if (lean_is_scalar(x_157)) { + x_159 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_151); - x_158 = lean_box(0); + x_159 = x_157; } -x_159 = lean_ctor_get(x_156, 1); -lean_inc(x_159); -lean_dec(x_156); -if (lean_is_scalar(x_158)) { +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_156); +if (lean_is_scalar(x_154)) { x_160 = lean_alloc_ctor(0, 2, 0); } else { - x_160 = x_158; + x_160 = x_154; } lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_157); -if (lean_is_scalar(x_155)) { +lean_ctor_set(x_160, 1, x_153); +if (lean_is_scalar(x_152)) { x_161 = lean_alloc_ctor(0, 2, 0); } else { - x_161 = x_155; + x_161 = x_152; } lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_154); -if (lean_is_scalar(x_153)) { - x_162 = lean_alloc_ctor(0, 2, 0); -} else { - x_162 = x_153; -} -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_152); -return x_162; +lean_ctor_set(x_161, 1, x_151); +return x_161; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_163 = lean_ctor_get(x_149, 1); -lean_inc(x_163); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_162 = lean_ctor_get(x_148, 1); +lean_inc(x_162); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_163 = x_148; +} else { + lean_dec_ref(x_148); + x_163 = lean_box(0); +} +x_164 = lean_ctor_get(x_149, 1); +lean_inc(x_164); if (lean_is_exclusive(x_149)) { lean_ctor_release(x_149, 0); lean_ctor_release(x_149, 1); - x_164 = x_149; + x_165 = x_149; } else { lean_dec_ref(x_149); - x_164 = lean_box(0); + x_165 = lean_box(0); } -x_165 = lean_ctor_get(x_150, 1); -lean_inc(x_165); +x_166 = lean_ctor_get(x_150, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_150, 1); +lean_inc(x_167); if (lean_is_exclusive(x_150)) { lean_ctor_release(x_150, 0); lean_ctor_release(x_150, 1); - x_166 = x_150; + x_168 = x_150; } else { lean_dec_ref(x_150); - x_166 = lean_box(0); + x_168 = lean_box(0); } -x_167 = lean_ctor_get(x_151, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_151, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_169 = x_151; +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_151); - x_169 = lean_box(0); + x_169 = x_168; } -if (lean_is_scalar(x_169)) { - x_170 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +if (lean_is_scalar(x_165)) { + x_170 = lean_alloc_ctor(0, 2, 0); } else { - x_170 = x_169; + x_170 = x_165; } -lean_ctor_set(x_170, 0, x_167); -lean_ctor_set(x_170, 1, x_168); -if (lean_is_scalar(x_166)) { +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_164); +if (lean_is_scalar(x_163)) { x_171 = lean_alloc_ctor(0, 2, 0); -} else { - x_171 = x_166; -} -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_165); -if (lean_is_scalar(x_164)) { - x_172 = lean_alloc_ctor(0, 2, 0); -} else { - x_172 = x_164; -} -lean_ctor_set(x_172, 0, x_171); -lean_ctor_set(x_172, 1, x_163); -return x_172; -} -} -else -{ -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_173 = lean_ctor_get(x_149, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_149, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_175 = x_149; -} else { - lean_dec_ref(x_149); - x_175 = lean_box(0); -} -if (lean_is_scalar(x_175)) { - x_176 = lean_alloc_ctor(1, 2, 0); -} else { - x_176 = x_175; -} -lean_ctor_set(x_176, 0, x_173); -lean_ctor_set(x_176, 1, x_174); -return x_176; -} -} -} -} -} -else -{ -uint8_t x_177; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_177 = !lean_is_exclusive(x_9); -if (x_177 == 0) -{ -lean_object* x_178; uint8_t x_179; -x_178 = lean_ctor_get(x_9, 0); -lean_dec(x_178); -x_179 = !lean_is_exclusive(x_11); -if (x_179 == 0) -{ -lean_object* x_180; -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_9); -lean_ctor_set(x_180, 1, x_10); -return x_180; -} -else -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_181 = lean_ctor_get(x_11, 0); -x_182 = lean_ctor_get(x_11, 1); -lean_inc(x_182); -lean_inc(x_181); -lean_dec(x_11); -x_183 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_183, 0, x_181); -lean_ctor_set(x_183, 1, x_182); -lean_ctor_set(x_9, 0, x_183); -x_184 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_184, 0, x_9); -lean_ctor_set(x_184, 1, x_10); -return x_184; -} -} -else -{ -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_185 = lean_ctor_get(x_9, 1); -lean_inc(x_185); -lean_dec(x_9); -x_186 = lean_ctor_get(x_11, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_11, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - lean_ctor_release(x_11, 1); - x_188 = x_11; -} else { - lean_dec_ref(x_11); - x_188 = lean_box(0); -} -if (lean_is_scalar(x_188)) { - x_189 = lean_alloc_ctor(1, 2, 0); -} else { - x_189 = x_188; -} -lean_ctor_set(x_189, 0, x_186); -lean_ctor_set(x_189, 1, x_187); -x_190 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_190, 0, x_189); -lean_ctor_set(x_190, 1, x_185); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_190); -lean_ctor_set(x_191, 1, x_10); -return x_191; -} -} -} -block_270: -{ -lean_object* x_207; -x_207 = lean_ctor_get(x_205, 0); -lean_inc(x_207); -if (lean_obj_tag(x_207) == 0) -{ -uint8_t x_208; -x_208 = !lean_is_exclusive(x_205); -if (x_208 == 0) -{ -lean_object* x_209; uint8_t x_210; -x_209 = lean_ctor_get(x_205, 0); -lean_dec(x_209); -x_210 = !lean_is_exclusive(x_207); -if (x_210 == 0) -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_207, 0); -x_212 = lean_ctor_get(x_207, 1); -x_213 = l_Lean_parseImports_x27(x_211, x_204, x_206); -lean_dec(x_204); -lean_dec(x_211); -if (lean_obj_tag(x_213) == 0) -{ -lean_object* x_214; lean_object* x_215; -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); -lean_ctor_set(x_207, 0, x_214); -x_9 = x_205; -x_10 = x_215; -goto block_192; -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; uint8_t x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_216 = lean_ctor_get(x_213, 0); -lean_inc(x_216); -x_217 = lean_ctor_get(x_213, 1); -lean_inc(x_217); -lean_dec(x_213); -x_218 = lean_io_error_to_string(x_216); -x_219 = 3; -x_220 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set_uint8(x_220, sizeof(void*)*1, x_219); -x_221 = lean_array_get_size(x_212); -x_222 = lean_array_push(x_212, x_220); -lean_ctor_set_tag(x_207, 1); -lean_ctor_set(x_207, 1, x_222); -lean_ctor_set(x_207, 0, x_221); -x_9 = x_205; -x_10 = x_217; -goto block_192; -} -} -else -{ -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_207, 0); -x_224 = lean_ctor_get(x_207, 1); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_207); -x_225 = l_Lean_parseImports_x27(x_223, x_204, x_206); -lean_dec(x_204); -lean_dec(x_223); -if (lean_obj_tag(x_225) == 0) -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_226 = lean_ctor_get(x_225, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_225, 1); -lean_inc(x_227); -lean_dec(x_225); -x_228 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_228, 0, x_226); -lean_ctor_set(x_228, 1, x_224); -lean_ctor_set(x_205, 0, x_228); -x_9 = x_205; -x_10 = x_227; -goto block_192; -} -else -{ -lean_object* x_229; lean_object* x_230; lean_object* x_231; uint8_t x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_229 = lean_ctor_get(x_225, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_225, 1); -lean_inc(x_230); -lean_dec(x_225); -x_231 = lean_io_error_to_string(x_229); -x_232 = 3; -x_233 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_233, 0, x_231); -lean_ctor_set_uint8(x_233, sizeof(void*)*1, x_232); -x_234 = lean_array_get_size(x_224); -x_235 = lean_array_push(x_224, x_233); -x_236 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_236, 0, x_234); -lean_ctor_set(x_236, 1, x_235); -lean_ctor_set(x_205, 0, x_236); -x_9 = x_205; -x_10 = x_230; -goto block_192; -} -} -} -else -{ -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_237 = lean_ctor_get(x_205, 1); -lean_inc(x_237); -lean_dec(x_205); -x_238 = lean_ctor_get(x_207, 0); -lean_inc(x_238); -x_239 = lean_ctor_get(x_207, 1); -lean_inc(x_239); -if (lean_is_exclusive(x_207)) { - lean_ctor_release(x_207, 0); - lean_ctor_release(x_207, 1); - x_240 = x_207; -} else { - lean_dec_ref(x_207); - x_240 = lean_box(0); -} -x_241 = l_Lean_parseImports_x27(x_238, x_204, x_206); -lean_dec(x_204); -lean_dec(x_238); -if (lean_obj_tag(x_241) == 0) -{ -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_242 = lean_ctor_get(x_241, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -lean_dec(x_241); -if (lean_is_scalar(x_240)) { - x_244 = lean_alloc_ctor(0, 2, 0); -} else { - x_244 = x_240; -} -lean_ctor_set(x_244, 0, x_242); -lean_ctor_set(x_244, 1, x_239); -x_245 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_245, 0, x_244); -lean_ctor_set(x_245, 1, x_237); -x_9 = x_245; -x_10 = x_243; -goto block_192; -} -else -{ -lean_object* x_246; lean_object* x_247; lean_object* x_248; uint8_t x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_246 = lean_ctor_get(x_241, 0); -lean_inc(x_246); -x_247 = lean_ctor_get(x_241, 1); -lean_inc(x_247); -lean_dec(x_241); -x_248 = lean_io_error_to_string(x_246); -x_249 = 3; -x_250 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_250, 0, x_248); -lean_ctor_set_uint8(x_250, sizeof(void*)*1, x_249); -x_251 = lean_array_get_size(x_239); -x_252 = lean_array_push(x_239, x_250); -if (lean_is_scalar(x_240)) { - x_253 = lean_alloc_ctor(1, 2, 0); -} else { - x_253 = x_240; - lean_ctor_set_tag(x_253, 1); -} -lean_ctor_set(x_253, 0, x_251); -lean_ctor_set(x_253, 1, x_252); -x_254 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_254, 0, x_253); -lean_ctor_set(x_254, 1, x_237); -x_9 = x_254; -x_10 = x_247; -goto block_192; -} -} -} -else -{ -uint8_t x_255; -lean_dec(x_204); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_255 = !lean_is_exclusive(x_205); -if (x_255 == 0) -{ -lean_object* x_256; uint8_t x_257; -x_256 = lean_ctor_get(x_205, 0); -lean_dec(x_256); -x_257 = !lean_is_exclusive(x_207); -if (x_257 == 0) -{ -lean_object* x_258; -x_258 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_258, 0, x_205); -lean_ctor_set(x_258, 1, x_206); -return x_258; -} -else -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_259 = lean_ctor_get(x_207, 0); -x_260 = lean_ctor_get(x_207, 1); -lean_inc(x_260); -lean_inc(x_259); -lean_dec(x_207); -x_261 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_205, 0, x_261); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_205); -lean_ctor_set(x_262, 1, x_206); -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; -x_263 = lean_ctor_get(x_205, 1); -lean_inc(x_263); -lean_dec(x_205); -x_264 = lean_ctor_get(x_207, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_207, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_207)) { - lean_ctor_release(x_207, 0); - lean_ctor_release(x_207, 1); - x_266 = x_207; -} else { - lean_dec_ref(x_207); - x_266 = lean_box(0); -} -if (lean_is_scalar(x_266)) { - x_267 = lean_alloc_ctor(1, 2, 0); -} else { - x_267 = x_266; -} -lean_ctor_set(x_267, 0, x_264); -lean_ctor_set(x_267, 1, x_265); -x_268 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_263); -x_269 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_269, 0, x_268); -lean_ctor_set(x_269, 1, x_206); -return x_269; -} -} -} -} -} -static lean_object* _init_l_Lake_Module_recParseImports___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("module source file not found: ", 30, 30); -return x_1; -} -} -static lean_object* _init_l_Lake_Module_recParseImports___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lake_Module_recParseImports(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_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_108 = lean_ctor_get(x_1, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 2); -lean_inc(x_111); -lean_dec(x_109); -x_112 = lean_ctor_get(x_111, 7); -lean_inc(x_112); -lean_dec(x_111); -x_113 = l_System_FilePath_join(x_110, x_112); -x_114 = lean_ctor_get(x_108, 1); -lean_inc(x_114); -lean_dec(x_108); -x_115 = lean_ctor_get(x_114, 2); -lean_inc(x_115); -lean_dec(x_114); -x_116 = l_System_FilePath_join(x_113, x_115); -x_117 = lean_ctor_get(x_1, 1); -lean_inc(x_117); -x_118 = l_Lake_Module_recParseImports___lambda__1___closed__1; -x_119 = l_Lean_moduleExists_go(x_116, x_118, x_117, x_7); -lean_dec(x_116); -if (lean_obj_tag(x_119) == 0) -{ -uint8_t x_120; -x_120 = !lean_is_exclusive(x_119); -if (x_120 == 0) -{ -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_119, 1); -lean_ctor_set(x_119, 1, x_5); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_119); -lean_ctor_set(x_122, 1, x_6); -x_8 = x_122; -x_9 = x_121; -goto block_107; +} else { + x_171 = x_163; } -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_123 = lean_ctor_get(x_119, 0); -x_124 = lean_ctor_get(x_119, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_119); -x_125 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_5); -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_6); -x_8 = x_126; -x_9 = x_124; -goto block_107; +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_162); +return x_171; } } else { -uint8_t x_127; -x_127 = !lean_is_exclusive(x_119); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_128 = lean_ctor_get(x_119, 0); -x_129 = lean_ctor_get(x_119, 1); -x_130 = lean_io_error_to_string(x_128); -x_131 = 3; -x_132 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set_uint8(x_132, sizeof(void*)*1, x_131); -x_133 = lean_array_get_size(x_5); -x_134 = lean_array_push(x_5, x_132); -lean_ctor_set(x_119, 1, x_134); -lean_ctor_set(x_119, 0, x_133); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_119); -lean_ctor_set(x_135, 1, x_6); -x_8 = x_135; -x_9 = x_129; -goto block_107; -} -else -{ -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; -x_136 = lean_ctor_get(x_119, 0); -x_137 = lean_ctor_get(x_119, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_119); -x_138 = lean_io_error_to_string(x_136); -x_139 = 3; -x_140 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set_uint8(x_140, sizeof(void*)*1, x_139); -x_141 = lean_array_get_size(x_5); -x_142 = lean_array_push(x_5, x_140); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_6); -x_8 = x_144; -x_9 = x_137; -goto block_107; +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_172 = lean_ctor_get(x_148, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_148, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_174 = x_148; +} else { + lean_dec_ref(x_148); + x_174 = lean_box(0); } +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(1, 2, 0); +} else { + x_175 = x_174; } -block_107: -{ -lean_object* x_10; -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; +} +} +} +} +} +else { -uint8_t x_13; +uint8_t x_176; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) +x_176 = !lean_is_exclusive(x_8); +if (x_176 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_8, 0); -lean_dec(x_14); -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) +lean_object* x_177; uint8_t x_178; +x_177 = lean_ctor_get(x_8, 0); +lean_dec(x_177); +x_178 = !lean_is_exclusive(x_10); +if (x_178 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_16 = lean_ctor_get(x_10, 1); -x_17 = lean_ctor_get(x_10, 0); -lean_dec(x_17); -x_18 = lean_ctor_get(x_1, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 2); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_21, 7); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_System_FilePath_join(x_20, x_22); -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); -lean_dec(x_18); -x_25 = lean_ctor_get(x_24, 2); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_System_FilePath_join(x_23, x_25); -x_27 = lean_ctor_get(x_1, 1); -lean_inc(x_27); -lean_dec(x_1); -x_28 = l_Lake_Module_recParseImports___lambda__1___closed__1; -x_29 = l_Lean_modToFilePath(x_26, x_27, x_28); -lean_dec(x_26); -x_30 = l_Lake_Module_recParseImports___closed__1; -x_31 = lean_string_append(x_30, x_29); -lean_dec(x_29); -x_32 = l_Lake_Module_recParseImports___closed__2; -x_33 = lean_string_append(x_31, x_32); -x_34 = 3; -x_35 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_34); -x_36 = lean_array_get_size(x_16); -x_37 = lean_array_push(x_16, x_35); -lean_ctor_set_tag(x_10, 1); -lean_ctor_set(x_10, 1, x_37); -lean_ctor_set(x_10, 0, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_8); -lean_ctor_set(x_38, 1, x_9); -return x_38; +lean_object* x_179; +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_8); +lean_ctor_set(x_179, 1, x_9); +return x_179; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_39 = lean_ctor_get(x_10, 1); -lean_inc(x_39); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_180 = lean_ctor_get(x_10, 0); +x_181 = lean_ctor_get(x_10, 1); +lean_inc(x_181); +lean_inc(x_180); lean_dec(x_10); -x_40 = lean_ctor_get(x_1, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 2); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_ctor_get(x_43, 7); -lean_inc(x_44); -lean_dec(x_43); -x_45 = l_System_FilePath_join(x_42, x_44); -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_dec(x_40); -x_47 = lean_ctor_get(x_46, 2); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l_System_FilePath_join(x_45, x_47); -x_49 = lean_ctor_get(x_1, 1); -lean_inc(x_49); -lean_dec(x_1); -x_50 = l_Lake_Module_recParseImports___lambda__1___closed__1; -x_51 = l_Lean_modToFilePath(x_48, x_49, x_50); -lean_dec(x_48); -x_52 = l_Lake_Module_recParseImports___closed__1; -x_53 = lean_string_append(x_52, x_51); -lean_dec(x_51); -x_54 = l_Lake_Module_recParseImports___closed__2; -x_55 = lean_string_append(x_53, x_54); -x_56 = 3; -x_57 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set_uint8(x_57, sizeof(void*)*1, x_56); -x_58 = lean_array_get_size(x_39); -x_59 = lean_array_push(x_39, x_57); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_8, 0, x_60); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_8); -lean_ctor_set(x_61, 1, x_9); -return x_61; +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_8, 0, x_182); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_8); +lean_ctor_set(x_183, 1, x_9); +return x_183; } } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_62 = lean_ctor_get(x_8, 1); -lean_inc(x_62); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_184 = lean_ctor_get(x_8, 1); +lean_inc(x_184); lean_dec(x_8); -x_63 = lean_ctor_get(x_10, 1); -lean_inc(x_63); +x_185 = lean_ctor_get(x_10, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_10, 1); +lean_inc(x_186); if (lean_is_exclusive(x_10)) { lean_ctor_release(x_10, 0); lean_ctor_release(x_10, 1); - x_64 = x_10; + x_187 = x_10; } else { lean_dec_ref(x_10); - x_64 = lean_box(0); + x_187 = lean_box(0); } -x_65 = lean_ctor_get(x_1, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 2); -lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_68, 7); -lean_inc(x_69); -lean_dec(x_68); -x_70 = l_System_FilePath_join(x_67, x_69); -x_71 = lean_ctor_get(x_65, 1); -lean_inc(x_71); -lean_dec(x_65); -x_72 = lean_ctor_get(x_71, 2); -lean_inc(x_72); -lean_dec(x_71); -x_73 = l_System_FilePath_join(x_70, x_72); -x_74 = lean_ctor_get(x_1, 1); -lean_inc(x_74); -lean_dec(x_1); -x_75 = l_Lake_Module_recParseImports___lambda__1___closed__1; -x_76 = l_Lean_modToFilePath(x_73, x_74, x_75); -lean_dec(x_73); -x_77 = l_Lake_Module_recParseImports___closed__1; -x_78 = lean_string_append(x_77, x_76); -lean_dec(x_76); -x_79 = l_Lake_Module_recParseImports___closed__2; -x_80 = lean_string_append(x_78, x_79); -x_81 = 3; -x_82 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set_uint8(x_82, sizeof(void*)*1, x_81); -x_83 = lean_array_get_size(x_63); -x_84 = lean_array_push(x_63, x_82); -if (lean_is_scalar(x_64)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_187)) { + x_188 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_64; - lean_ctor_set_tag(x_85, 1); + x_188 = x_187; } -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_62); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_9); -return x_87; +lean_ctor_set(x_188, 0, x_185); +lean_ctor_set(x_188, 1, x_186); +x_189 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_184); +x_190 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_190, 0, x_189); +lean_ctor_set(x_190, 1, x_9); +return x_190; } } +} +block_269: +{ +lean_object* x_206; +x_206 = lean_ctor_get(x_204, 0); +lean_inc(x_206); +if (lean_obj_tag(x_206) == 0) +{ +uint8_t x_207; +x_207 = !lean_is_exclusive(x_204); +if (x_207 == 0) +{ +lean_object* x_208; uint8_t x_209; +x_208 = lean_ctor_get(x_204, 0); +lean_dec(x_208); +x_209 = !lean_is_exclusive(x_206); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_210 = lean_ctor_get(x_206, 0); +x_211 = lean_ctor_get(x_206, 1); +x_212 = l_Lean_parseImports_x27(x_210, x_203, x_205); +lean_dec(x_203); +lean_dec(x_210); +if (lean_obj_tag(x_212) == 0) +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_ctor_get(x_212, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_212, 1); +lean_inc(x_214); +lean_dec(x_212); +lean_ctor_set(x_206, 0, x_213); +x_8 = x_204; +x_9 = x_214; +goto block_191; +} else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_88 = lean_ctor_get(x_8, 1); -lean_inc(x_88); -lean_dec(x_8); -x_89 = lean_ctor_get(x_10, 1); -lean_inc(x_89); -lean_dec(x_10); -x_90 = lean_box(0); -x_91 = l_Lake_Module_recParseImports___lambda__1(x_1, x_90, x_2, x_3, x_4, x_89, x_88, x_9); -return x_91; +lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_215 = lean_ctor_get(x_212, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_212, 1); +lean_inc(x_216); +lean_dec(x_212); +x_217 = lean_io_error_to_string(x_215); +x_218 = 3; +x_219 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_219, 0, x_217); +lean_ctor_set_uint8(x_219, sizeof(void*)*1, x_218); +x_220 = lean_array_get_size(x_211); +x_221 = lean_array_push(x_211, x_219); +lean_ctor_set_tag(x_206, 1); +lean_ctor_set(x_206, 1, x_221); +lean_ctor_set(x_206, 0, x_220); +x_8 = x_204; +x_9 = x_216; +goto block_191; +} +} +else +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_206, 0); +x_223 = lean_ctor_get(x_206, 1); +lean_inc(x_223); +lean_inc(x_222); +lean_dec(x_206); +x_224 = l_Lean_parseImports_x27(x_222, x_203, x_205); +lean_dec(x_203); +lean_dec(x_222); +if (lean_obj_tag(x_224) == 0) +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_224, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_224, 1); +lean_inc(x_226); +lean_dec(x_224); +x_227 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_223); +lean_ctor_set(x_204, 0, x_227); +x_8 = x_204; +x_9 = x_226; +goto block_191; +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_228 = lean_ctor_get(x_224, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_224, 1); +lean_inc(x_229); +lean_dec(x_224); +x_230 = lean_io_error_to_string(x_228); +x_231 = 3; +x_232 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_232, 0, x_230); +lean_ctor_set_uint8(x_232, sizeof(void*)*1, x_231); +x_233 = lean_array_get_size(x_223); +x_234 = lean_array_push(x_223, x_232); +x_235 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_235, 0, x_233); +lean_ctor_set(x_235, 1, x_234); +lean_ctor_set(x_204, 0, x_235); +x_8 = x_204; +x_9 = x_229; +goto block_191; +} } } else { -uint8_t x_92; +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +x_236 = lean_ctor_get(x_204, 1); +lean_inc(x_236); +lean_dec(x_204); +x_237 = lean_ctor_get(x_206, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_206, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_239 = x_206; +} else { + lean_dec_ref(x_206); + x_239 = lean_box(0); +} +x_240 = l_Lean_parseImports_x27(x_237, x_203, x_205); +lean_dec(x_203); +lean_dec(x_237); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +x_241 = lean_ctor_get(x_240, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_240, 1); +lean_inc(x_242); +lean_dec(x_240); +if (lean_is_scalar(x_239)) { + x_243 = lean_alloc_ctor(0, 2, 0); +} else { + x_243 = x_239; +} +lean_ctor_set(x_243, 0, x_241); +lean_ctor_set(x_243, 1, x_238); +x_244 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_236); +x_8 = x_244; +x_9 = x_242; +goto block_191; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_245 = lean_ctor_get(x_240, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_240, 1); +lean_inc(x_246); +lean_dec(x_240); +x_247 = lean_io_error_to_string(x_245); +x_248 = 3; +x_249 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_249, 0, x_247); +lean_ctor_set_uint8(x_249, sizeof(void*)*1, x_248); +x_250 = lean_array_get_size(x_238); +x_251 = lean_array_push(x_238, x_249); +if (lean_is_scalar(x_239)) { + x_252 = lean_alloc_ctor(1, 2, 0); +} else { + x_252 = x_239; + lean_ctor_set_tag(x_252, 1); +} +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +x_253 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_253, 0, x_252); +lean_ctor_set(x_253, 1, x_236); +x_8 = x_253; +x_9 = x_246; +goto block_191; +} +} +} +else +{ +uint8_t x_254; +lean_dec(x_203); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_92 = !lean_is_exclusive(x_8); -if (x_92 == 0) +x_254 = !lean_is_exclusive(x_204); +if (x_254 == 0) { -lean_object* x_93; uint8_t x_94; -x_93 = lean_ctor_get(x_8, 0); -lean_dec(x_93); -x_94 = !lean_is_exclusive(x_10); -if (x_94 == 0) +lean_object* x_255; uint8_t x_256; +x_255 = lean_ctor_get(x_204, 0); +lean_dec(x_255); +x_256 = !lean_is_exclusive(x_206); +if (x_256 == 0) { -lean_object* x_95; -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_8); -lean_ctor_set(x_95, 1, x_9); -return x_95; +lean_object* x_257; +x_257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_257, 0, x_204); +lean_ctor_set(x_257, 1, x_205); +return x_257; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_96 = lean_ctor_get(x_10, 0); -x_97 = lean_ctor_get(x_10, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_10); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_8, 0, x_98); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_8); -lean_ctor_set(x_99, 1, x_9); -return x_99; +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_258 = lean_ctor_get(x_206, 0); +x_259 = lean_ctor_get(x_206, 1); +lean_inc(x_259); +lean_inc(x_258); +lean_dec(x_206); +x_260 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +lean_ctor_set(x_204, 0, x_260); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_204); +lean_ctor_set(x_261, 1, x_205); +return x_261; } } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_100 = lean_ctor_get(x_8, 1); -lean_inc(x_100); -lean_dec(x_8); -x_101 = lean_ctor_get(x_10, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_10, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - x_103 = x_10; +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_262 = lean_ctor_get(x_204, 1); +lean_inc(x_262); +lean_dec(x_204); +x_263 = lean_ctor_get(x_206, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_206, 1); +lean_inc(x_264); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_265 = x_206; } else { - lean_dec_ref(x_10); - x_103 = lean_box(0); + lean_dec_ref(x_206); + x_265 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_265)) { + x_266 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_266 = x_265; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_100); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_9); -return x_106; +lean_ctor_set(x_266, 0, x_263); +lean_ctor_set(x_266, 1, x_264); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_266); +lean_ctor_set(x_267, 1, x_262); +x_268 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_268, 0, x_267); +lean_ctor_set(x_268, 1, x_205); +return x_268; } } } @@ -4628,15 +4203,6 @@ lean_dec(x_1); return x_13; } } -LEAN_EXPORT lean_object* l_Lake_Module_recParseImports___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lake_Module_recParseImports___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -return x_9; -} -} LEAN_EXPORT lean_object* l_Lake_Module_importsFacetConfig___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { @@ -4657,7 +4223,7 @@ static lean_object* _init_l_Lake_Module_importsFacetConfig___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_1 = l_Lake_Module_recParseImports___closed__1; x_2 = l_Lake_Module_importsFacetConfig___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4750,7 +4316,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lake_collectImportsAux__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(": bad import '", 14, 14); +x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } @@ -4758,6 +4324,14 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lake_collectImportsAux__ _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked(": bad import '", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("'", 1, 1); return x_1; } @@ -5160,9 +4734,9 @@ x_119 = lean_ctor_get(x_111, 0); x_120 = lean_ctor_get(x_111, 1); x_121 = l_Array_shrink___rarg(x_120, x_119); lean_dec(x_119); -x_122 = l_Lake_Module_recParseImports___closed__2; +x_122 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_123 = lean_string_append(x_122, x_1); -x_124 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_124 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_125 = lean_string_append(x_123, x_124); x_126 = lean_ctor_get(x_105, 1); lean_inc(x_126); @@ -5171,7 +4745,7 @@ x_127 = 1; x_128 = l_Lean_Name_toString(x_126, x_127); x_129 = lean_string_append(x_125, x_128); lean_dec(x_128); -x_130 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_130 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_131 = lean_string_append(x_129, x_130); x_132 = 3; x_133 = lean_alloc_ctor(0, 1, 1); @@ -5207,9 +4781,9 @@ lean_inc(x_139); lean_dec(x_111); x_141 = l_Array_shrink___rarg(x_140, x_139); lean_dec(x_139); -x_142 = l_Lake_Module_recParseImports___closed__2; +x_142 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_143 = lean_string_append(x_142, x_1); -x_144 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_144 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_145 = lean_string_append(x_143, x_144); x_146 = lean_ctor_get(x_105, 1); lean_inc(x_146); @@ -5218,7 +4792,7 @@ x_147 = 1; x_148 = l_Lean_Name_toString(x_146, x_147); x_149 = lean_string_append(x_145, x_148); lean_dec(x_148); -x_150 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_150 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_151 = lean_string_append(x_149, x_150); x_152 = 3; x_153 = lean_alloc_ctor(0, 1, 1); @@ -5266,9 +4840,9 @@ if (lean_is_exclusive(x_111)) { } x_164 = l_Array_shrink___rarg(x_162, x_161); lean_dec(x_161); -x_165 = l_Lake_Module_recParseImports___closed__2; +x_165 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_166 = lean_string_append(x_165, x_1); -x_167 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_167 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_168 = lean_string_append(x_166, x_167); x_169 = lean_ctor_get(x_105, 1); lean_inc(x_169); @@ -5277,7 +4851,7 @@ x_170 = 1; x_171 = l_Lean_Name_toString(x_169, x_170); x_172 = lean_string_append(x_168, x_171); lean_dec(x_171); -x_173 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_173 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_174 = lean_string_append(x_172, x_173); x_175 = 3; x_176 = lean_alloc_ctor(0, 1, 1); @@ -6175,7 +5749,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lake_Module_recComputeTr _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_1 = l_Lake_Module_recParseImports___closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lake_Module_recComputeTransImports___spec__1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -6477,9 +6051,9 @@ x_118 = lean_ctor_get(x_110, 0); x_119 = lean_ctor_get(x_110, 1); x_120 = l_Array_shrink___rarg(x_119, x_118); lean_dec(x_118); -x_121 = l_Lake_Module_recParseImports___closed__2; +x_121 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_122 = lean_string_append(x_121, x_1); -x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_124 = lean_string_append(x_122, x_123); x_125 = lean_ctor_get(x_104, 1); lean_inc(x_125); @@ -6488,7 +6062,7 @@ x_126 = 1; x_127 = l_Lean_Name_toString(x_125, x_126); x_128 = lean_string_append(x_124, x_127); lean_dec(x_127); -x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_130 = lean_string_append(x_128, x_129); x_131 = 3; x_132 = lean_alloc_ctor(0, 1, 1); @@ -6524,9 +6098,9 @@ lean_inc(x_138); lean_dec(x_110); x_140 = l_Array_shrink___rarg(x_139, x_138); lean_dec(x_138); -x_141 = l_Lake_Module_recParseImports___closed__2; +x_141 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_142 = lean_string_append(x_141, x_1); -x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_144 = lean_string_append(x_142, x_143); x_145 = lean_ctor_get(x_104, 1); lean_inc(x_145); @@ -6535,7 +6109,7 @@ x_146 = 1; x_147 = l_Lean_Name_toString(x_145, x_146); x_148 = lean_string_append(x_144, x_147); lean_dec(x_147); -x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_150 = lean_string_append(x_148, x_149); x_151 = 3; x_152 = lean_alloc_ctor(0, 1, 1); @@ -6583,9 +6157,9 @@ if (lean_is_exclusive(x_110)) { } x_163 = l_Array_shrink___rarg(x_161, x_160); lean_dec(x_160); -x_164 = l_Lake_Module_recParseImports___closed__2; +x_164 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_165 = lean_string_append(x_164, x_1); -x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_167 = lean_string_append(x_165, x_166); x_168 = lean_ctor_get(x_104, 1); lean_inc(x_168); @@ -6594,7 +6168,7 @@ x_169 = 1; x_170 = l_Lean_Name_toString(x_168, x_169); x_171 = lean_string_append(x_167, x_170); lean_dec(x_170); -x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_173 = lean_string_append(x_171, x_172); x_174 = 3; x_175 = lean_alloc_ctor(0, 1, 1); @@ -7357,7 +6931,7 @@ x_26 = l_System_FilePath_join(x_23, x_25); x_27 = lean_ctor_get(x_1, 1); lean_inc(x_27); lean_dec(x_1); -x_28 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_28 = l_Lake_Module_recParseImports___closed__1; x_29 = l_Lean_modToFilePath(x_26, x_27, x_28); lean_dec(x_26); x_30 = lean_array_size(x_16); @@ -7977,7 +7551,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lake_computePrecompileIm _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_1 = l_Lake_Module_recParseImports___closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lake_computePrecompileImportsAux___spec__1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -8722,9 +8296,9 @@ x_118 = lean_ctor_get(x_110, 0); x_119 = lean_ctor_get(x_110, 1); x_120 = l_Array_shrink___rarg(x_119, x_118); lean_dec(x_118); -x_121 = l_Lake_Module_recParseImports___closed__2; +x_121 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_122 = lean_string_append(x_121, x_1); -x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_124 = lean_string_append(x_122, x_123); x_125 = lean_ctor_get(x_104, 1); lean_inc(x_125); @@ -8733,7 +8307,7 @@ x_126 = 1; x_127 = l_Lean_Name_toString(x_125, x_126); x_128 = lean_string_append(x_124, x_127); lean_dec(x_127); -x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_130 = lean_string_append(x_128, x_129); x_131 = 3; x_132 = lean_alloc_ctor(0, 1, 1); @@ -8769,9 +8343,9 @@ lean_inc(x_138); lean_dec(x_110); x_140 = l_Array_shrink___rarg(x_139, x_138); lean_dec(x_138); -x_141 = l_Lake_Module_recParseImports___closed__2; +x_141 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_142 = lean_string_append(x_141, x_1); -x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_144 = lean_string_append(x_142, x_143); x_145 = lean_ctor_get(x_104, 1); lean_inc(x_145); @@ -8780,7 +8354,7 @@ x_146 = 1; x_147 = l_Lean_Name_toString(x_145, x_146); x_148 = lean_string_append(x_144, x_147); lean_dec(x_147); -x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_150 = lean_string_append(x_148, x_149); x_151 = 3; x_152 = lean_alloc_ctor(0, 1, 1); @@ -8828,9 +8402,9 @@ if (lean_is_exclusive(x_110)) { } x_163 = l_Array_shrink___rarg(x_161, x_160); lean_dec(x_160); -x_164 = l_Lake_Module_recParseImports___closed__2; +x_164 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_165 = lean_string_append(x_164, x_1); -x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_167 = lean_string_append(x_165, x_166); x_168 = lean_ctor_get(x_104, 1); lean_inc(x_168); @@ -8839,7 +8413,7 @@ x_169 = 1; x_170 = l_Lean_Name_toString(x_168, x_169); x_171 = lean_string_append(x_167, x_170); lean_dec(x_170); -x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_173 = lean_string_append(x_171, x_172); x_174 = 3; x_175 = lean_alloc_ctor(0, 1, 1); @@ -10690,9 +10264,9 @@ x_118 = lean_ctor_get(x_110, 0); x_119 = lean_ctor_get(x_110, 1); x_120 = l_Array_shrink___rarg(x_119, x_118); lean_dec(x_118); -x_121 = l_Lake_Module_recParseImports___closed__2; +x_121 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_122 = lean_string_append(x_121, x_1); -x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_123 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_124 = lean_string_append(x_122, x_123); x_125 = lean_ctor_get(x_104, 1); lean_inc(x_125); @@ -10701,7 +10275,7 @@ x_126 = 1; x_127 = l_Lean_Name_toString(x_125, x_126); x_128 = lean_string_append(x_124, x_127); lean_dec(x_127); -x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_129 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_130 = lean_string_append(x_128, x_129); x_131 = 3; x_132 = lean_alloc_ctor(0, 1, 1); @@ -10737,9 +10311,9 @@ lean_inc(x_138); lean_dec(x_110); x_140 = l_Array_shrink___rarg(x_139, x_138); lean_dec(x_138); -x_141 = l_Lake_Module_recParseImports___closed__2; +x_141 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_142 = lean_string_append(x_141, x_1); -x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_143 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_144 = lean_string_append(x_142, x_143); x_145 = lean_ctor_get(x_104, 1); lean_inc(x_145); @@ -10748,7 +10322,7 @@ x_146 = 1; x_147 = l_Lean_Name_toString(x_145, x_146); x_148 = lean_string_append(x_144, x_147); lean_dec(x_147); -x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_149 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_150 = lean_string_append(x_148, x_149); x_151 = 3; x_152 = lean_alloc_ctor(0, 1, 1); @@ -10796,9 +10370,9 @@ if (lean_is_exclusive(x_110)) { } x_163 = l_Array_shrink___rarg(x_161, x_160); lean_dec(x_160); -x_164 = l_Lake_Module_recParseImports___closed__2; +x_164 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_165 = lean_string_append(x_164, x_1); -x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; +x_166 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; x_167 = lean_string_append(x_165, x_166); x_168 = lean_ctor_get(x_104, 1); lean_inc(x_168); @@ -10807,7 +10381,7 @@ x_169 = 1; x_170 = l_Lean_Name_toString(x_168, x_169); x_171 = lean_string_append(x_167, x_170); lean_dec(x_170); -x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2; +x_172 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3; x_173 = lean_string_append(x_171, x_172); x_174 = 3; x_175 = lean_alloc_ctor(0, 1, 1); @@ -11570,7 +11144,7 @@ x_26 = l_System_FilePath_join(x_23, x_25); x_27 = lean_ctor_get(x_1, 1); lean_inc(x_27); lean_dec(x_1); -x_28 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_28 = l_Lake_Module_recParseImports___closed__1; x_29 = l_Lean_modToFilePath(x_26, x_27, x_28); lean_dec(x_26); x_30 = lean_array_size(x_16); @@ -12463,10 +12037,10 @@ x_67 = lean_ctor_get(x_66, 2); lean_inc(x_67); lean_dec(x_66); x_68 = l_System_FilePath_join(x_65, x_67); -x_69 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_69 = l_Lake_Module_recParseImports___closed__1; x_70 = l_Lean_modToFilePath(x_68, x_20, x_69); lean_dec(x_68); -x_71 = l_Lake_Module_recParseImports___closed__2; +x_71 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_72 = lean_string_append(x_71, x_70); lean_dec(x_70); x_73 = l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDeps___spec__1___closed__1; @@ -24082,7 +23656,7 @@ x_37 = lean_ctor_get(x_23, 2); lean_inc(x_37); lean_dec(x_23); x_38 = l_System_FilePath_join(x_36, x_37); -x_39 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_39 = l_Lake_Module_recParseImports___closed__1; lean_inc(x_2); x_40 = l_Lean_modToFilePath(x_38, x_2, x_39); lean_inc(x_40); @@ -24136,7 +23710,7 @@ lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint x_53 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_54 = lean_string_append(x_53, x_50); lean_dec(x_50); -x_55 = l_Lake_Module_recParseImports___closed__2; +x_55 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_56 = lean_string_append(x_54, x_55); x_57 = 1; x_58 = lean_alloc_ctor(0, 1, 1); @@ -24324,7 +23898,7 @@ x_115 = lean_ctor_get(x_100, 2); lean_inc(x_115); lean_dec(x_100); x_116 = l_System_FilePath_join(x_114, x_115); -x_117 = l_Lake_Module_recParseImports___lambda__1___closed__1; +x_117 = l_Lake_Module_recParseImports___closed__1; lean_inc(x_2); x_118 = l_Lean_modToFilePath(x_116, x_2, x_117); lean_inc(x_118); @@ -24378,7 +23952,7 @@ lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; x_131 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_132 = lean_string_append(x_131, x_128); lean_dec(x_128); -x_133 = l_Lake_Module_recParseImports___closed__2; +x_133 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_134 = lean_string_append(x_132, x_133); x_135 = 1; x_136 = lean_alloc_ctor(0, 1, 1); @@ -25050,7 +24624,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint x_31 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_32 = lean_string_append(x_31, x_28); lean_dec(x_28); -x_33 = l_Lake_Module_recParseImports___closed__2; +x_33 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_34 = lean_string_append(x_32, x_33); x_35 = 1; x_36 = lean_alloc_ctor(0, 1, 1); @@ -25910,7 +25484,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint x_31 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_32 = lean_string_append(x_31, x_28); lean_dec(x_28); -x_33 = l_Lake_Module_recParseImports___closed__2; +x_33 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_34 = lean_string_append(x_32, x_33); x_35 = 1; x_36 = lean_alloc_ctor(0, 1, 1); @@ -26726,7 +26300,7 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint x_29 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_30 = lean_string_append(x_29, x_26); lean_dec(x_26); -x_31 = l_Lake_Module_recParseImports___closed__2; +x_31 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_32 = lean_string_append(x_30, x_31); x_33 = 1; x_34 = lean_alloc_ctor(0, 1, 1); @@ -27535,7 +27109,7 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint x_29 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_30 = lean_string_append(x_29, x_26); lean_dec(x_26); -x_31 = l_Lake_Module_recParseImports___closed__2; +x_31 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_32 = lean_string_append(x_30, x_31); x_33 = 1; x_34 = lean_alloc_ctor(0, 1, 1); @@ -28813,7 +28387,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint x_23 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_20); lean_dec(x_20); -x_25 = l_Lake_Module_recParseImports___closed__2; +x_25 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_26 = lean_string_append(x_24, x_25); x_27 = 1; x_28 = lean_alloc_ctor(0, 1, 1); @@ -29246,7 +28820,7 @@ lean_inc(x_12); x_13 = 1; lean_inc(x_12); x_14 = l_Lean_Name_toString(x_12, x_13); -x_15 = l_Lake_Module_recParseImports___closed__2; +x_15 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_16 = lean_string_append(x_15, x_14); lean_dec(x_14); x_17 = l_Lake_Module_recBuildLeanCToOExport___closed__1; @@ -29685,7 +29259,7 @@ lean_inc(x_12); x_13 = 1; lean_inc(x_12); x_14 = l_Lean_Name_toString(x_12, x_13); -x_15 = l_Lake_Module_recParseImports___closed__2; +x_15 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_16 = lean_string_append(x_15, x_14); lean_dec(x_14); x_17 = l_Lake_Module_recBuildLeanCToOExport___closed__1; @@ -30078,7 +29652,7 @@ lean_inc(x_8); x_9 = 1; lean_inc(x_8); x_10 = l_Lean_Name_toString(x_8, x_9); -x_11 = l_Lake_Module_recParseImports___closed__2; +x_11 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); x_13 = l_Lake_Module_recBuildLeanBcToO___closed__1; @@ -30714,7 +30288,7 @@ x_7 = lean_array_uset(x_3, x_2, x_6); x_8 = l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDynlib___spec__3___closed__1; x_9 = lean_string_append(x_8, x_5); lean_dec(x_5); -x_10 = l_Lake_Module_recParseImports___closed__2; +x_10 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_11 = lean_string_append(x_9, x_10); x_12 = 1; x_13 = lean_usize_add(x_2, x_12); @@ -30751,7 +30325,7 @@ x_7 = lean_array_uset(x_3, x_2, x_6); x_8 = l_Array_mapMUnsafe_map___at_Lake_Module_recBuildDynlib___spec__4___closed__1; x_9 = lean_string_append(x_8, x_5); lean_dec(x_5); -x_10 = l_Lake_Module_recParseImports___closed__2; +x_10 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_11 = lean_string_append(x_9, x_10); x_12 = 1; x_13 = lean_usize_add(x_2, x_12); @@ -31339,7 +30913,7 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint x_39 = l_Lake_Module_recBuildLean___lambda__5___closed__1; x_40 = lean_string_append(x_39, x_36); lean_dec(x_36); -x_41 = l_Lake_Module_recParseImports___closed__2; +x_41 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_42 = lean_string_append(x_40, x_41); x_43 = 1; x_44 = lean_alloc_ctor(0, 1, 1); @@ -32803,7 +32377,7 @@ lean_inc(x_8); x_9 = 1; lean_inc(x_8); x_10 = l_Lean_Name_toString(x_8, x_9); -x_11 = l_Lake_Module_recParseImports___closed__2; +x_11 = l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); x_13 = l_Lake_Module_recBuildDynlib___closed__1; @@ -33243,12 +32817,8 @@ l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__ lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__2); l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__3(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lake_Module_recParseImports___spec__12___closed__3); -l_Lake_Module_recParseImports___lambda__1___closed__1 = _init_l_Lake_Module_recParseImports___lambda__1___closed__1(); -lean_mark_persistent(l_Lake_Module_recParseImports___lambda__1___closed__1); l_Lake_Module_recParseImports___closed__1 = _init_l_Lake_Module_recParseImports___closed__1(); lean_mark_persistent(l_Lake_Module_recParseImports___closed__1); -l_Lake_Module_recParseImports___closed__2 = _init_l_Lake_Module_recParseImports___closed__2(); -lean_mark_persistent(l_Lake_Module_recParseImports___closed__2); l_Lake_Module_importsFacetConfig___closed__1 = _init_l_Lake_Module_importsFacetConfig___closed__1(); lean_mark_persistent(l_Lake_Module_importsFacetConfig___closed__1); l_Lake_Module_importsFacetConfig___closed__2 = _init_l_Lake_Module_importsFacetConfig___closed__2(); @@ -33263,6 +32833,8 @@ l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1 = _in lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__1); l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__2); +l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lake_collectImportsAux___spec__3___closed__3); l_Lake_collectImportsAux___closed__1 = _init_l_Lake_collectImportsAux___closed__1(); lean_mark_persistent(l_Lake_collectImportsAux___closed__1); l_Array_forInUnsafe_loop___at_Lake_Module_recComputeTransImports___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lake_Module_recComputeTransImports___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Compiler/CSimpAttr.c b/stage0/stdlib/Lean/Compiler/CSimpAttr.c index cf3609c7e544..ce238e4d671e 100644 --- a/stage0/stdlib/Lean/Compiler/CSimpAttr.c +++ b/stage0/stdlib/Lean/Compiler/CSimpAttr.c @@ -34,6 +34,7 @@ uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_474____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Compiler_CSimp_add___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_138_(lean_object*); +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_hasCSimpAttribute___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_138____closed__4; size_t lean_uint64_to_usize(uint64_t); @@ -109,7 +110,6 @@ size_t lean_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT uint8_t l_Lean_Compiler_hasCSimpAttribute(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Compiler_CSimp_add___spec__2___closed__1; -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_138____closed__5; lean_object* l_Lean_Expr_constLevels_x21(lean_object*); lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -350,7 +350,7 @@ lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_3); -x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(x_4); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); return x_1; @@ -364,7 +364,7 @@ lean_inc(x_8); lean_inc(x_7); lean_dec(x_1); x_9 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(x_8); x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index b4fa4f1c531e..4cfaa5324b04 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -241,6 +241,7 @@ lean_object* l_Lean_IR_EmitC_emitUSet(lean_object*, lean_object*, lean_object*, static lean_object* l_String_foldlAux___at_Lean_IR_EmitC_quoteString___spec__1___closed__4; LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_IR_EmitC_emitArgs___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_EmitC_overwriteParam(lean_object*, lean_object*); +static lean_object* l_Lean_IR_EmitC_shouldExport___closed__12; extern lean_object* l_Lean_exportAttr; LEAN_EXPORT lean_object* l_List_forM___at_Lean_IR_EmitC_emitLns___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitTailCall___closed__1; @@ -416,6 +417,7 @@ static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__30; static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__7; extern lean_object* l_Lean_NameSet_empty; static lean_object* l_List_forM___at_Lean_IR_EmitC_emitFnDecls___spec__5___closed__1; +static lean_object* l_Lean_IR_EmitC_shouldExport___closed__13; static lean_object* l_Lean_IR_EmitC_emitPartialApp___closed__2; static lean_object* l_Lean_IR_EmitC_emitDeclAux___lambda__3___closed__2; lean_object* l_Lean_IR_EmitC_emitCtorScalarSize(lean_object*, lean_object*, lean_object*, lean_object*); @@ -597,6 +599,7 @@ static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__7; lean_object* l_Lean_IR_collectUsedDecls(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__1___closed__15; static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__16; +static lean_object* l_Lean_IR_EmitC_shouldExport___closed__11; 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*); @@ -1850,28 +1853,56 @@ return x_1; static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Watchdog", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_IR_EmitC_shouldExport___closed__1; x_2 = l_Lean_IR_EmitC_shouldExport___closed__7; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +x_3 = l_Lean_IR_EmitC_shouldExport___closed__8; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__9() { +static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("findModuleRefs", 14, 14); +x_1 = lean_mk_string_unchecked("ImportCompletion", 16, 16); return x_1; } } -static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__10() { +static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_IR_EmitC_shouldExport___closed__1; +x_2 = l_Lean_IR_EmitC_shouldExport___closed__7; +x_3 = l_Lean_IR_EmitC_shouldExport___closed__10; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Completion", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_IR_EmitC_shouldExport___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_IR_EmitC_shouldExport___closed__1; x_2 = l_Lean_IR_EmitC_shouldExport___closed__7; -x_3 = l_Lean_IR_EmitC_shouldExport___closed__9; +x_3 = l_Lean_IR_EmitC_shouldExport___closed__12; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } @@ -1890,34 +1921,57 @@ x_5 = l_Lean_Name_isPrefixOf(x_4, x_1); if (x_5 == 0) { lean_object* x_6; uint8_t x_7; -x_6 = l_Lean_IR_EmitC_shouldExport___closed__8; +x_6 = l_Lean_IR_EmitC_shouldExport___closed__9; x_7 = l_Lean_Name_isPrefixOf(x_6, x_1); if (x_7 == 0) { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_8; uint8_t x_9; +x_8 = l_Lean_IR_EmitC_shouldExport___closed__11; +x_9 = l_Lean_Name_isPrefixOf(x_8, x_1); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_IR_EmitC_shouldExport___closed__13; +x_11 = l_Lean_Name_isPrefixOf(x_10, x_1); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = 1; +return x_12; } else { -lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_IR_EmitC_shouldExport___closed__10; -x_10 = lean_name_eq(x_1, x_9); -return x_10; +uint8_t x_13; +x_13 = 0; +return x_13; } } else { -uint8_t x_11; -x_11 = 0; -return x_11; +uint8_t x_14; +x_14 = 0; +return x_14; } } else { -uint8_t x_12; -x_12 = 0; -return x_12; +uint8_t x_15; +x_15 = 0; +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = 0; +return x_17; } } } @@ -15752,6 +15806,12 @@ l_Lean_IR_EmitC_shouldExport___closed__9 = _init_l_Lean_IR_EmitC_shouldExport___ lean_mark_persistent(l_Lean_IR_EmitC_shouldExport___closed__9); l_Lean_IR_EmitC_shouldExport___closed__10 = _init_l_Lean_IR_EmitC_shouldExport___closed__10(); lean_mark_persistent(l_Lean_IR_EmitC_shouldExport___closed__10); +l_Lean_IR_EmitC_shouldExport___closed__11 = _init_l_Lean_IR_EmitC_shouldExport___closed__11(); +lean_mark_persistent(l_Lean_IR_EmitC_shouldExport___closed__11); +l_Lean_IR_EmitC_shouldExport___closed__12 = _init_l_Lean_IR_EmitC_shouldExport___closed__12(); +lean_mark_persistent(l_Lean_IR_EmitC_shouldExport___closed__12); +l_Lean_IR_EmitC_shouldExport___closed__13 = _init_l_Lean_IR_EmitC_shouldExport___closed__13(); +lean_mark_persistent(l_Lean_IR_EmitC_shouldExport___closed__13); l_Nat_forM_loop___at_Lean_IR_EmitC_emitFnDeclAux___spec__1___closed__1 = _init_l_Nat_forM_loop___at_Lean_IR_EmitC_emitFnDeclAux___spec__1___closed__1(); lean_mark_persistent(l_Nat_forM_loop___at_Lean_IR_EmitC_emitFnDeclAux___spec__1___closed__1); l_Lean_IR_EmitC_emitFnDeclAux___lambda__1___closed__1 = _init_l_Lean_IR_EmitC_emitFnDeclAux___lambda__1___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index 9c8160fd85be..9dc3195b9269 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -225,6 +225,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen_declRange__1___clo static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabInitQuot_declRange__1___closed__1; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce_go(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange__1___closed__7; @@ -315,6 +316,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabExport___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabExport___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange__1___closed__2; @@ -367,6 +369,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabExit__1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabExport___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheck___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc_declRange__1___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2; LEAN_EXPORT lean_object* l_Lean_activateScoped___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSection_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabModuleDoc___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -404,6 +407,7 @@ lean_object* l_Lean_MessageData_signature(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption_declRange__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAddDeclDoc__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth__1___closed__1; +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__3; static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__19___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(lean_object*); @@ -412,6 +416,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabOpe LEAN_EXPORT lean_object* l_List_map___at_Lean_Elab_Command_elabExport___spec__16(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__4; size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getRefPos___at_Lean_Elab_Command_elabExport___spec__20___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabInitQuot_declRange__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalCoreUnsafe___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -473,6 +478,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabEval__1___closed__2; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSection___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabExit__1___closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption_declRange__1___closed__6; static lean_object* l_Lean_Elab_Command_elabNamespace___closed__1; static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__13; @@ -496,6 +502,7 @@ static lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___closed__4; static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at_Lean_Elab_Command_elabExport___spec__18___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabVariable___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__5; +lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__4___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange__1___closed__1; lean_object* l_Lean_ScopedEnvExtension_popScope___rarg(lean_object*, lean_object*); @@ -579,8 +586,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange__1__ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabOpen___spec__2(lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Command_elabOpen___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabExit(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption__1___closed__3; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_activateScoped___at_Lean_Elab_Command_elabOpen___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -590,6 +599,7 @@ static lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveName static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange__1___closed__7; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -657,6 +667,7 @@ static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_ad lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__5; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabExport___spec__11___closed__3; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabInclude__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalBang(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabVariable___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection_declRange__1___closed__7; @@ -720,6 +731,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__5___boxed(l static lean_object* l_Lean_Elab_Command_elabEvalCoreUnsafe___lambda__3___closed__46; extern lean_object* l_Lean_scopedEnvExtensionsRef; static lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___closed__12; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAddDeclDoc___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabImport___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(lean_object*); @@ -750,6 +762,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabRunMeta_declRange_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabVariable___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_add_alias(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabRunElab_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth__1___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -803,6 +816,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Comma static lean_object* l_Lean_Elab_Command_elabSection___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabRunCmd_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasNoErrorMessages___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabRunElab_declRange__1___closed__6; static lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___closed__3; @@ -849,6 +863,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange__1___cl LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabImport___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabExport_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd__1___closed__5; +static lean_object* l_Lean_Elab_Command_elabInclude___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange__1___closed__2; @@ -1027,6 +1042,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabRunCmd_declRange__1___c static lean_object* l_Lean_Elab_Command_elabEvalCoreUnsafe___lambda__3___closed__44; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckFailure(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabAddDeclDoc___closed__6; +static lean_object* l_Lean_Elab_Command_elabInclude___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveId___at_Lean_Elab_Command_elabExport___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabEnd___closed__1; LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Command_elabOpen___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1116,6 +1132,7 @@ lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange__1___closed__2; lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__7; +LEAN_EXPORT lean_object* l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__4___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabEvalCoreUnsafe___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Command_elabExport___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1124,6 +1141,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCom static lean_object* l___regBuiltin_Lean_Elab_Command_elabAddDeclDoc_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Lean_getRefPos___at_Lean_Elab_Command_elabExport___spec__20(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabRunElab__1(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at_Lean_Elab_Command_elabExport___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2294,7 +2312,7 @@ lean_inc(x_4); x_26 = l_Lean_Environment_registerNamespace(x_24, x_4); x_27 = l_Lean_Elab_Command_instInhabitedScope; x_28 = l_List_head_x21___rarg(x_27, x_25); -x_29 = lean_ctor_get_uint8(x_28, sizeof(void*)*7); +x_29 = lean_ctor_get_uint8(x_28, sizeof(void*)*8); if (x_29 == 0) { uint8_t x_30; @@ -2309,7 +2327,7 @@ lean_dec(x_32); lean_inc(x_4); lean_ctor_set(x_28, 2, x_4); lean_ctor_set(x_28, 0, x_3); -lean_ctor_set_uint8(x_28, sizeof(void*)*7, x_2); +lean_ctor_set_uint8(x_28, sizeof(void*)*8, x_2); lean_ctor_set_tag(x_19, 1); lean_ctor_set(x_19, 1, x_25); lean_ctor_set(x_19, 0, x_28); @@ -2324,12 +2342,14 @@ goto block_18; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; x_35 = lean_ctor_get(x_28, 1); x_36 = lean_ctor_get(x_28, 3); x_37 = lean_ctor_get(x_28, 4); x_38 = lean_ctor_get(x_28, 5); x_39 = lean_ctor_get(x_28, 6); +x_40 = lean_ctor_get(x_28, 7); +lean_inc(x_40); lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); @@ -2337,410 +2357,430 @@ lean_inc(x_36); lean_inc(x_35); lean_dec(x_28); lean_inc(x_4); -x_40 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_40, 0, x_3); -lean_ctor_set(x_40, 1, x_35); -lean_ctor_set(x_40, 2, x_4); -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 4, x_37); -lean_ctor_set(x_40, 5, x_38); -lean_ctor_set(x_40, 6, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*7, x_2); +x_41 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_35); +lean_ctor_set(x_41, 2, x_4); +lean_ctor_set(x_41, 3, x_36); +lean_ctor_set(x_41, 4, x_37); +lean_ctor_set(x_41, 5, x_38); +lean_ctor_set(x_41, 6, x_39); +lean_ctor_set(x_41, 7, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*8, x_2); lean_ctor_set_tag(x_19, 1); lean_ctor_set(x_19, 1, x_25); -lean_ctor_set(x_19, 0, x_40); +lean_ctor_set(x_19, 0, x_41); lean_ctor_set(x_21, 2, x_19); lean_ctor_set(x_21, 0, x_26); -x_41 = lean_st_ref_set(x_6, x_21, x_23); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -x_8 = x_42; +x_42 = lean_st_ref_set(x_6, x_21, x_23); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_8 = x_43; goto block_18; } } else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_28); -if (x_43 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_28); +if (x_44 == 0) { -lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -x_44 = lean_ctor_get(x_28, 2); -lean_dec(x_44); -x_45 = lean_ctor_get(x_28, 0); +lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; +x_45 = lean_ctor_get(x_28, 2); lean_dec(x_45); -x_46 = 1; +x_46 = lean_ctor_get(x_28, 0); +lean_dec(x_46); +x_47 = 1; lean_inc(x_4); lean_ctor_set(x_28, 2, x_4); lean_ctor_set(x_28, 0, x_3); -lean_ctor_set_uint8(x_28, sizeof(void*)*7, x_46); +lean_ctor_set_uint8(x_28, sizeof(void*)*8, x_47); lean_ctor_set_tag(x_19, 1); lean_ctor_set(x_19, 1, x_25); lean_ctor_set(x_19, 0, x_28); lean_ctor_set(x_21, 2, x_19); lean_ctor_set(x_21, 0, x_26); -x_47 = lean_st_ref_set(x_6, x_21, x_23); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_8 = x_48; +x_48 = lean_st_ref_set(x_6, x_21, x_23); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_8 = x_49; goto block_18; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_49 = lean_ctor_get(x_28, 1); -x_50 = lean_ctor_get(x_28, 3); -x_51 = lean_ctor_get(x_28, 4); -x_52 = lean_ctor_get(x_28, 5); -x_53 = lean_ctor_get(x_28, 6); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_50 = lean_ctor_get(x_28, 1); +x_51 = lean_ctor_get(x_28, 3); +x_52 = lean_ctor_get(x_28, 4); +x_53 = lean_ctor_get(x_28, 5); +x_54 = lean_ctor_get(x_28, 6); +x_55 = lean_ctor_get(x_28, 7); +lean_inc(x_55); +lean_inc(x_54); lean_inc(x_53); lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); lean_dec(x_28); -x_54 = 1; +x_56 = 1; lean_inc(x_4); -x_55 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_55, 0, x_3); -lean_ctor_set(x_55, 1, x_49); -lean_ctor_set(x_55, 2, x_4); -lean_ctor_set(x_55, 3, x_50); -lean_ctor_set(x_55, 4, x_51); -lean_ctor_set(x_55, 5, x_52); -lean_ctor_set(x_55, 6, x_53); -lean_ctor_set_uint8(x_55, sizeof(void*)*7, x_54); +x_57 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_57, 0, x_3); +lean_ctor_set(x_57, 1, x_50); +lean_ctor_set(x_57, 2, x_4); +lean_ctor_set(x_57, 3, x_51); +lean_ctor_set(x_57, 4, x_52); +lean_ctor_set(x_57, 5, x_53); +lean_ctor_set(x_57, 6, x_54); +lean_ctor_set(x_57, 7, x_55); +lean_ctor_set_uint8(x_57, sizeof(void*)*8, x_56); lean_ctor_set_tag(x_19, 1); lean_ctor_set(x_19, 1, x_25); -lean_ctor_set(x_19, 0, x_55); +lean_ctor_set(x_19, 0, x_57); lean_ctor_set(x_21, 2, x_19); lean_ctor_set(x_21, 0, x_26); -x_56 = lean_st_ref_set(x_6, x_21, x_23); -x_57 = lean_ctor_get(x_56, 1); -lean_inc(x_57); -lean_dec(x_56); -x_8 = x_57; +x_58 = lean_st_ref_set(x_6, x_21, x_23); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_8 = x_59; goto block_18; } } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_58 = lean_ctor_get(x_19, 1); -x_59 = lean_ctor_get(x_21, 0); -x_60 = lean_ctor_get(x_21, 1); -x_61 = lean_ctor_get(x_21, 2); -x_62 = lean_ctor_get(x_21, 3); -x_63 = lean_ctor_get(x_21, 4); -x_64 = lean_ctor_get(x_21, 5); -x_65 = lean_ctor_get(x_21, 6); -x_66 = lean_ctor_get(x_21, 7); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_60 = lean_ctor_get(x_19, 1); +x_61 = lean_ctor_get(x_21, 0); +x_62 = lean_ctor_get(x_21, 1); +x_63 = lean_ctor_get(x_21, 2); +x_64 = lean_ctor_get(x_21, 3); +x_65 = lean_ctor_get(x_21, 4); +x_66 = lean_ctor_get(x_21, 5); +x_67 = lean_ctor_get(x_21, 6); +x_68 = lean_ctor_get(x_21, 7); +lean_inc(x_68); +lean_inc(x_67); lean_inc(x_66); lean_inc(x_65); lean_inc(x_64); lean_inc(x_63); lean_inc(x_62); lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); lean_dec(x_21); lean_inc(x_4); -x_67 = l_Lean_Environment_registerNamespace(x_59, x_4); -x_68 = l_Lean_Elab_Command_instInhabitedScope; -x_69 = l_List_head_x21___rarg(x_68, x_61); -x_70 = lean_ctor_get_uint8(x_69, sizeof(void*)*7); -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; -x_71 = lean_ctor_get(x_69, 1); -lean_inc(x_71); -x_72 = lean_ctor_get(x_69, 3); -lean_inc(x_72); -x_73 = lean_ctor_get(x_69, 4); +x_69 = l_Lean_Environment_registerNamespace(x_61, x_4); +x_70 = l_Lean_Elab_Command_instInhabitedScope; +x_71 = l_List_head_x21___rarg(x_70, x_63); +x_72 = lean_ctor_get_uint8(x_71, sizeof(void*)*8); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); -x_74 = lean_ctor_get(x_69, 5); +x_74 = lean_ctor_get(x_71, 3); lean_inc(x_74); -x_75 = lean_ctor_get(x_69, 6); +x_75 = lean_ctor_get(x_71, 4); lean_inc(x_75); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - lean_ctor_release(x_69, 2); - lean_ctor_release(x_69, 3); - lean_ctor_release(x_69, 4); - lean_ctor_release(x_69, 5); - lean_ctor_release(x_69, 6); - x_76 = x_69; -} else { - lean_dec_ref(x_69); - x_76 = lean_box(0); +x_76 = lean_ctor_get(x_71, 5); +lean_inc(x_76); +x_77 = lean_ctor_get(x_71, 6); +lean_inc(x_77); +x_78 = lean_ctor_get(x_71, 7); +lean_inc(x_78); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + lean_ctor_release(x_71, 2); + lean_ctor_release(x_71, 3); + lean_ctor_release(x_71, 4); + lean_ctor_release(x_71, 5); + lean_ctor_release(x_71, 6); + lean_ctor_release(x_71, 7); + x_79 = x_71; +} else { + lean_dec_ref(x_71); + x_79 = lean_box(0); } lean_inc(x_4); -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(0, 7, 1); -} else { - x_77 = x_76; -} -lean_ctor_set(x_77, 0, x_3); -lean_ctor_set(x_77, 1, x_71); -lean_ctor_set(x_77, 2, x_4); -lean_ctor_set(x_77, 3, x_72); -lean_ctor_set(x_77, 4, x_73); -lean_ctor_set(x_77, 5, x_74); -lean_ctor_set(x_77, 6, x_75); -lean_ctor_set_uint8(x_77, sizeof(void*)*7, x_2); +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(0, 8, 1); +} else { + x_80 = x_79; +} +lean_ctor_set(x_80, 0, x_3); +lean_ctor_set(x_80, 1, x_73); +lean_ctor_set(x_80, 2, x_4); +lean_ctor_set(x_80, 3, x_74); +lean_ctor_set(x_80, 4, x_75); +lean_ctor_set(x_80, 5, x_76); +lean_ctor_set(x_80, 6, x_77); +lean_ctor_set(x_80, 7, x_78); +lean_ctor_set_uint8(x_80, sizeof(void*)*8, x_2); lean_ctor_set_tag(x_19, 1); -lean_ctor_set(x_19, 1, x_61); -lean_ctor_set(x_19, 0, x_77); -x_78 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_78, 0, x_67); -lean_ctor_set(x_78, 1, x_60); -lean_ctor_set(x_78, 2, x_19); -lean_ctor_set(x_78, 3, x_62); -lean_ctor_set(x_78, 4, x_63); -lean_ctor_set(x_78, 5, x_64); -lean_ctor_set(x_78, 6, x_65); -lean_ctor_set(x_78, 7, x_66); -x_79 = lean_st_ref_set(x_6, x_78, x_58); -x_80 = lean_ctor_get(x_79, 1); -lean_inc(x_80); -lean_dec(x_79); -x_8 = x_80; +lean_ctor_set(x_19, 1, x_63); +lean_ctor_set(x_19, 0, x_80); +x_81 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_81, 0, x_69); +lean_ctor_set(x_81, 1, x_62); +lean_ctor_set(x_81, 2, x_19); +lean_ctor_set(x_81, 3, x_64); +lean_ctor_set(x_81, 4, x_65); +lean_ctor_set(x_81, 5, x_66); +lean_ctor_set(x_81, 6, x_67); +lean_ctor_set(x_81, 7, x_68); +x_82 = lean_st_ref_set(x_6, x_81, x_60); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +x_8 = x_83; goto block_18; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_81 = lean_ctor_get(x_69, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_69, 3); -lean_inc(x_82); -x_83 = lean_ctor_get(x_69, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_69, 5); +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; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_84 = lean_ctor_get(x_71, 1); lean_inc(x_84); -x_85 = lean_ctor_get(x_69, 6); +x_85 = lean_ctor_get(x_71, 3); lean_inc(x_85); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - lean_ctor_release(x_69, 2); - lean_ctor_release(x_69, 3); - lean_ctor_release(x_69, 4); - lean_ctor_release(x_69, 5); - lean_ctor_release(x_69, 6); - x_86 = x_69; -} else { - lean_dec_ref(x_69); - x_86 = lean_box(0); -} -x_87 = 1; +x_86 = lean_ctor_get(x_71, 4); +lean_inc(x_86); +x_87 = lean_ctor_get(x_71, 5); +lean_inc(x_87); +x_88 = lean_ctor_get(x_71, 6); +lean_inc(x_88); +x_89 = lean_ctor_get(x_71, 7); +lean_inc(x_89); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + lean_ctor_release(x_71, 2); + lean_ctor_release(x_71, 3); + lean_ctor_release(x_71, 4); + lean_ctor_release(x_71, 5); + lean_ctor_release(x_71, 6); + lean_ctor_release(x_71, 7); + x_90 = x_71; +} else { + lean_dec_ref(x_71); + x_90 = lean_box(0); +} +x_91 = 1; lean_inc(x_4); -if (lean_is_scalar(x_86)) { - x_88 = lean_alloc_ctor(0, 7, 1); -} else { - x_88 = x_86; -} -lean_ctor_set(x_88, 0, x_3); -lean_ctor_set(x_88, 1, x_81); -lean_ctor_set(x_88, 2, x_4); -lean_ctor_set(x_88, 3, x_82); -lean_ctor_set(x_88, 4, x_83); -lean_ctor_set(x_88, 5, x_84); -lean_ctor_set(x_88, 6, x_85); -lean_ctor_set_uint8(x_88, sizeof(void*)*7, x_87); +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(0, 8, 1); +} else { + x_92 = x_90; +} +lean_ctor_set(x_92, 0, x_3); +lean_ctor_set(x_92, 1, x_84); +lean_ctor_set(x_92, 2, x_4); +lean_ctor_set(x_92, 3, x_85); +lean_ctor_set(x_92, 4, x_86); +lean_ctor_set(x_92, 5, x_87); +lean_ctor_set(x_92, 6, x_88); +lean_ctor_set(x_92, 7, x_89); +lean_ctor_set_uint8(x_92, sizeof(void*)*8, x_91); lean_ctor_set_tag(x_19, 1); -lean_ctor_set(x_19, 1, x_61); -lean_ctor_set(x_19, 0, x_88); -x_89 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_89, 0, x_67); -lean_ctor_set(x_89, 1, x_60); -lean_ctor_set(x_89, 2, x_19); -lean_ctor_set(x_89, 3, x_62); -lean_ctor_set(x_89, 4, x_63); -lean_ctor_set(x_89, 5, x_64); -lean_ctor_set(x_89, 6, x_65); -lean_ctor_set(x_89, 7, x_66); -x_90 = lean_st_ref_set(x_6, x_89, x_58); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_8 = x_91; +lean_ctor_set(x_19, 1, x_63); +lean_ctor_set(x_19, 0, x_92); +x_93 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_93, 0, x_69); +lean_ctor_set(x_93, 1, x_62); +lean_ctor_set(x_93, 2, x_19); +lean_ctor_set(x_93, 3, x_64); +lean_ctor_set(x_93, 4, x_65); +lean_ctor_set(x_93, 5, x_66); +lean_ctor_set(x_93, 6, x_67); +lean_ctor_set(x_93, 7, x_68); +x_94 = lean_st_ref_set(x_6, x_93, x_60); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_8 = x_95; goto block_18; } } } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_92 = lean_ctor_get(x_19, 0); -x_93 = lean_ctor_get(x_19, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_19); -x_94 = lean_ctor_get(x_92, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_92, 1); -lean_inc(x_95); -x_96 = lean_ctor_get(x_92, 2); -lean_inc(x_96); -x_97 = lean_ctor_get(x_92, 3); +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; uint8_t x_110; +x_96 = lean_ctor_get(x_19, 0); +x_97 = lean_ctor_get(x_19, 1); lean_inc(x_97); -x_98 = lean_ctor_get(x_92, 4); +lean_inc(x_96); +lean_dec(x_19); +x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); -x_99 = lean_ctor_get(x_92, 5); +x_99 = lean_ctor_get(x_96, 1); lean_inc(x_99); -x_100 = lean_ctor_get(x_92, 6); +x_100 = lean_ctor_get(x_96, 2); lean_inc(x_100); -x_101 = lean_ctor_get(x_92, 7); +x_101 = lean_ctor_get(x_96, 3); lean_inc(x_101); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - lean_ctor_release(x_92, 2); - lean_ctor_release(x_92, 3); - lean_ctor_release(x_92, 4); - lean_ctor_release(x_92, 5); - lean_ctor_release(x_92, 6); - lean_ctor_release(x_92, 7); - x_102 = x_92; -} else { - lean_dec_ref(x_92); - x_102 = lean_box(0); +x_102 = lean_ctor_get(x_96, 4); +lean_inc(x_102); +x_103 = lean_ctor_get(x_96, 5); +lean_inc(x_103); +x_104 = lean_ctor_get(x_96, 6); +lean_inc(x_104); +x_105 = lean_ctor_get(x_96, 7); +lean_inc(x_105); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + lean_ctor_release(x_96, 2); + lean_ctor_release(x_96, 3); + lean_ctor_release(x_96, 4); + lean_ctor_release(x_96, 5); + lean_ctor_release(x_96, 6); + lean_ctor_release(x_96, 7); + x_106 = x_96; +} else { + lean_dec_ref(x_96); + x_106 = lean_box(0); } lean_inc(x_4); -x_103 = l_Lean_Environment_registerNamespace(x_94, x_4); -x_104 = l_Lean_Elab_Command_instInhabitedScope; -x_105 = l_List_head_x21___rarg(x_104, x_96); -x_106 = lean_ctor_get_uint8(x_105, sizeof(void*)*7); -if (x_106 == 0) -{ -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -x_108 = lean_ctor_get(x_105, 3); -lean_inc(x_108); -x_109 = lean_ctor_get(x_105, 4); -lean_inc(x_109); -x_110 = lean_ctor_get(x_105, 5); -lean_inc(x_110); -x_111 = lean_ctor_get(x_105, 6); +x_107 = l_Lean_Environment_registerNamespace(x_98, x_4); +x_108 = l_Lean_Elab_Command_instInhabitedScope; +x_109 = l_List_head_x21___rarg(x_108, x_100); +x_110 = lean_ctor_get_uint8(x_109, sizeof(void*)*8); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_111 = lean_ctor_get(x_109, 1); lean_inc(x_111); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - lean_ctor_release(x_105, 2); - lean_ctor_release(x_105, 3); - lean_ctor_release(x_105, 4); - lean_ctor_release(x_105, 5); - lean_ctor_release(x_105, 6); - x_112 = x_105; -} else { - lean_dec_ref(x_105); - x_112 = lean_box(0); +x_112 = lean_ctor_get(x_109, 3); +lean_inc(x_112); +x_113 = lean_ctor_get(x_109, 4); +lean_inc(x_113); +x_114 = lean_ctor_get(x_109, 5); +lean_inc(x_114); +x_115 = lean_ctor_get(x_109, 6); +lean_inc(x_115); +x_116 = lean_ctor_get(x_109, 7); +lean_inc(x_116); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + lean_ctor_release(x_109, 2); + lean_ctor_release(x_109, 3); + lean_ctor_release(x_109, 4); + lean_ctor_release(x_109, 5); + lean_ctor_release(x_109, 6); + lean_ctor_release(x_109, 7); + x_117 = x_109; +} else { + lean_dec_ref(x_109); + x_117 = lean_box(0); } lean_inc(x_4); -if (lean_is_scalar(x_112)) { - x_113 = lean_alloc_ctor(0, 7, 1); -} else { - x_113 = x_112; -} -lean_ctor_set(x_113, 0, x_3); -lean_ctor_set(x_113, 1, x_107); -lean_ctor_set(x_113, 2, x_4); -lean_ctor_set(x_113, 3, x_108); -lean_ctor_set(x_113, 4, x_109); -lean_ctor_set(x_113, 5, x_110); -lean_ctor_set(x_113, 6, x_111); -lean_ctor_set_uint8(x_113, sizeof(void*)*7, x_2); -x_114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_96); -if (lean_is_scalar(x_102)) { - x_115 = lean_alloc_ctor(0, 8, 0); -} else { - x_115 = x_102; -} -lean_ctor_set(x_115, 0, x_103); -lean_ctor_set(x_115, 1, x_95); -lean_ctor_set(x_115, 2, x_114); -lean_ctor_set(x_115, 3, x_97); -lean_ctor_set(x_115, 4, x_98); -lean_ctor_set(x_115, 5, x_99); -lean_ctor_set(x_115, 6, x_100); -lean_ctor_set(x_115, 7, x_101); -x_116 = lean_st_ref_set(x_6, x_115, x_93); -x_117 = lean_ctor_get(x_116, 1); -lean_inc(x_117); -lean_dec(x_116); -x_8 = x_117; +if (lean_is_scalar(x_117)) { + x_118 = lean_alloc_ctor(0, 8, 1); +} else { + x_118 = x_117; +} +lean_ctor_set(x_118, 0, x_3); +lean_ctor_set(x_118, 1, x_111); +lean_ctor_set(x_118, 2, x_4); +lean_ctor_set(x_118, 3, x_112); +lean_ctor_set(x_118, 4, x_113); +lean_ctor_set(x_118, 5, x_114); +lean_ctor_set(x_118, 6, x_115); +lean_ctor_set(x_118, 7, x_116); +lean_ctor_set_uint8(x_118, sizeof(void*)*8, x_2); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_100); +if (lean_is_scalar(x_106)) { + x_120 = lean_alloc_ctor(0, 8, 0); +} else { + x_120 = x_106; +} +lean_ctor_set(x_120, 0, x_107); +lean_ctor_set(x_120, 1, x_99); +lean_ctor_set(x_120, 2, x_119); +lean_ctor_set(x_120, 3, x_101); +lean_ctor_set(x_120, 4, x_102); +lean_ctor_set(x_120, 5, x_103); +lean_ctor_set(x_120, 6, x_104); +lean_ctor_set(x_120, 7, x_105); +x_121 = lean_st_ref_set(x_6, x_120, x_97); +x_122 = lean_ctor_get(x_121, 1); +lean_inc(x_122); +lean_dec(x_121); +x_8 = x_122; goto block_18; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_118 = lean_ctor_get(x_105, 1); -lean_inc(x_118); -x_119 = lean_ctor_get(x_105, 3); -lean_inc(x_119); -x_120 = lean_ctor_get(x_105, 4); -lean_inc(x_120); -x_121 = lean_ctor_get(x_105, 5); -lean_inc(x_121); -x_122 = lean_ctor_get(x_105, 6); -lean_inc(x_122); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - lean_ctor_release(x_105, 2); - lean_ctor_release(x_105, 3); - lean_ctor_release(x_105, 4); - lean_ctor_release(x_105, 5); - lean_ctor_release(x_105, 6); - x_123 = x_105; -} else { - lean_dec_ref(x_105); - x_123 = lean_box(0); +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_123 = lean_ctor_get(x_109, 1); +lean_inc(x_123); +x_124 = lean_ctor_get(x_109, 3); +lean_inc(x_124); +x_125 = lean_ctor_get(x_109, 4); +lean_inc(x_125); +x_126 = lean_ctor_get(x_109, 5); +lean_inc(x_126); +x_127 = lean_ctor_get(x_109, 6); +lean_inc(x_127); +x_128 = lean_ctor_get(x_109, 7); +lean_inc(x_128); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + lean_ctor_release(x_109, 2); + lean_ctor_release(x_109, 3); + lean_ctor_release(x_109, 4); + lean_ctor_release(x_109, 5); + lean_ctor_release(x_109, 6); + lean_ctor_release(x_109, 7); + x_129 = x_109; +} else { + lean_dec_ref(x_109); + x_129 = lean_box(0); } -x_124 = 1; +x_130 = 1; lean_inc(x_4); -if (lean_is_scalar(x_123)) { - x_125 = lean_alloc_ctor(0, 7, 1); -} else { - x_125 = x_123; -} -lean_ctor_set(x_125, 0, x_3); -lean_ctor_set(x_125, 1, x_118); -lean_ctor_set(x_125, 2, x_4); -lean_ctor_set(x_125, 3, x_119); -lean_ctor_set(x_125, 4, x_120); -lean_ctor_set(x_125, 5, x_121); -lean_ctor_set(x_125, 6, x_122); -lean_ctor_set_uint8(x_125, sizeof(void*)*7, x_124); -x_126 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_96); -if (lean_is_scalar(x_102)) { - x_127 = lean_alloc_ctor(0, 8, 0); -} else { - x_127 = x_102; -} -lean_ctor_set(x_127, 0, x_103); -lean_ctor_set(x_127, 1, x_95); -lean_ctor_set(x_127, 2, x_126); -lean_ctor_set(x_127, 3, x_97); -lean_ctor_set(x_127, 4, x_98); -lean_ctor_set(x_127, 5, x_99); -lean_ctor_set(x_127, 6, x_100); -lean_ctor_set(x_127, 7, x_101); -x_128 = lean_st_ref_set(x_6, x_127, x_93); -x_129 = lean_ctor_get(x_128, 1); -lean_inc(x_129); -lean_dec(x_128); -x_8 = x_129; +if (lean_is_scalar(x_129)) { + x_131 = lean_alloc_ctor(0, 8, 1); +} else { + x_131 = x_129; +} +lean_ctor_set(x_131, 0, x_3); +lean_ctor_set(x_131, 1, x_123); +lean_ctor_set(x_131, 2, x_4); +lean_ctor_set(x_131, 3, x_124); +lean_ctor_set(x_131, 4, x_125); +lean_ctor_set(x_131, 5, x_126); +lean_ctor_set(x_131, 6, x_127); +lean_ctor_set(x_131, 7, x_128); +lean_ctor_set_uint8(x_131, sizeof(void*)*8, x_130); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_100); +if (lean_is_scalar(x_106)) { + x_133 = lean_alloc_ctor(0, 8, 0); +} else { + x_133 = x_106; +} +lean_ctor_set(x_133, 0, x_107); +lean_ctor_set(x_133, 1, x_99); +lean_ctor_set(x_133, 2, x_132); +lean_ctor_set(x_133, 3, x_101); +lean_ctor_set(x_133, 4, x_102); +lean_ctor_set(x_133, 5, x_103); +lean_ctor_set(x_133, 6, x_104); +lean_ctor_set(x_133, 7, x_105); +x_134 = lean_st_ref_set(x_6, x_133, x_97); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +lean_dec(x_134); +x_8 = x_135; goto block_18; } } @@ -14709,14 +14749,16 @@ return x_2; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); x_7 = lean_ctor_get(x_2, 2); x_8 = lean_ctor_get(x_2, 4); x_9 = lean_ctor_get(x_2, 5); x_10 = lean_ctor_get(x_2, 6); -x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_11 = lean_ctor_get(x_2, 7); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -14724,16 +14766,17 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_6); -lean_ctor_set(x_12, 2, x_7); -lean_ctor_set(x_12, 3, x_1); -lean_ctor_set(x_12, 4, x_8); -lean_ctor_set(x_12, 5, x_9); -lean_ctor_set(x_12, 6, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*7, x_11); -return x_12; +x_13 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_6); +lean_ctor_set(x_13, 2, x_7); +lean_ctor_set(x_13, 3, x_1); +lean_ctor_set(x_13, 4, x_8); +lean_ctor_set(x_13, 5, x_9); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*8, x_12); +return x_13; } } } @@ -20078,14 +20121,16 @@ return x_2; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); x_7 = lean_ctor_get(x_2, 2); x_8 = lean_ctor_get(x_2, 3); x_9 = lean_ctor_get(x_2, 4); x_10 = lean_ctor_get(x_2, 6); -x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_11 = lean_ctor_get(x_2, 7); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -20093,16 +20138,17 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_6); -lean_ctor_set(x_12, 2, x_7); -lean_ctor_set(x_12, 3, x_8); -lean_ctor_set(x_12, 4, x_9); -lean_ctor_set(x_12, 5, x_1); -lean_ctor_set(x_12, 6, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*7, x_11); -return x_12; +x_13 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_6); +lean_ctor_set(x_13, 2, x_7); +lean_ctor_set(x_13, 3, x_8); +lean_ctor_set(x_13, 4, x_9); +lean_ctor_set(x_13, 5, x_1); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*8, x_12); +return x_13; } } } @@ -20754,16 +20800,16 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_5 = lean_ctor_get(x_3, 5); -x_6 = lean_ctor_get(x_3, 6); +x_6 = lean_ctor_get(x_3, 7); x_7 = lean_array_push(x_5, x_1); x_8 = l_Array_append___rarg(x_6, x_2); -lean_ctor_set(x_3, 6, x_8); +lean_ctor_set(x_3, 7, x_8); lean_ctor_set(x_3, 5, x_7); return x_3; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_9 = lean_ctor_get(x_3, 0); x_10 = lean_ctor_get(x_3, 1); x_11 = lean_ctor_get(x_3, 2); @@ -20771,7 +20817,9 @@ x_12 = lean_ctor_get(x_3, 3); x_13 = lean_ctor_get(x_3, 4); x_14 = lean_ctor_get(x_3, 5); x_15 = lean_ctor_get(x_3, 6); -x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*7); +x_16 = lean_ctor_get(x_3, 7); +x_17 = lean_ctor_get_uint8(x_3, sizeof(void*)*8); +lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -20780,18 +20828,19 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_dec(x_3); -x_17 = lean_array_push(x_14, x_1); -x_18 = l_Array_append___rarg(x_15, x_2); -x_19 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_19, 0, x_9); -lean_ctor_set(x_19, 1, x_10); -lean_ctor_set(x_19, 2, x_11); -lean_ctor_set(x_19, 3, x_12); -lean_ctor_set(x_19, 4, x_13); -lean_ctor_set(x_19, 5, x_17); -lean_ctor_set(x_19, 6, x_18); -lean_ctor_set_uint8(x_19, sizeof(void*)*7, x_16); -return x_19; +x_18 = lean_array_push(x_14, x_1); +x_19 = l_Array_append___rarg(x_16, x_2); +x_20 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_20, 0, x_9); +lean_ctor_set(x_20, 1, x_10); +lean_ctor_set(x_20, 2, x_11); +lean_ctor_set(x_20, 3, x_12); +lean_ctor_set(x_20, 4, x_13); +lean_ctor_set(x_20, 5, x_18); +lean_ctor_set(x_20, 6, x_15); +lean_ctor_set(x_20, 7, x_19); +lean_ctor_set_uint8(x_20, sizeof(void*)*8, x_17); +return x_20; } } } @@ -33439,14 +33488,16 @@ return x_2; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 2); x_7 = lean_ctor_get(x_2, 3); x_8 = lean_ctor_get(x_2, 4); x_9 = lean_ctor_get(x_2, 5); x_10 = lean_ctor_get(x_2, 6); -x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_11 = lean_ctor_get(x_2, 7); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -33454,16 +33505,17 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_1); -lean_ctor_set(x_12, 2, x_6); -lean_ctor_set(x_12, 3, x_7); -lean_ctor_set(x_12, 4, x_8); -lean_ctor_set(x_12, 5, x_9); -lean_ctor_set(x_12, 6, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*7, x_11); -return x_12; +x_13 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_1); +lean_ctor_set(x_13, 2, x_6); +lean_ctor_set(x_13, 3, x_7); +lean_ctor_set(x_13, 4, x_8); +lean_ctor_set(x_13, 5, x_9); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*8, x_12); +return x_13; } } } @@ -34928,6 +34980,407 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid 'include', variable '", 29, 29); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' has not been declared in the current scope", 44, 44); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_4, x_3); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_5); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_dec(x_5); +x_11 = lean_array_uget(x_2, x_4); +x_12 = l_Lean_Syntax_getId(x_11); +x_13 = l_Array_contains___at_Lean_findField_x3f___spec__1(x_1, x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_14 = l_Lean_MessageData_ofSyntax(x_11); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18(x_18, x_6, x_7, x_8); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +size_t x_24; size_t x_25; lean_object* x_26; +lean_dec(x_11); +x_24 = 1; +x_25 = lean_usize_add(x_4, x_24); +x_26 = lean_box(0); +x_4 = x_25; +x_5 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Lean_Syntax_getId(x_4); +lean_dec(x_4); +x_7 = l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(x_5); +lean_ctor_set(x_1, 1, x_7); +lean_ctor_set(x_1, 0, x_6); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = l_Lean_Syntax_getId(x_8); +lean_dec(x_8); +x_11 = l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 6); +x_5 = lean_array_to_list(lean_box(0), x_1); +x_6 = l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(x_5); +x_7 = l_List_appendTR___rarg(x_4, x_6); +lean_ctor_set(x_2, 6, x_7); +return x_2; +} +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; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_ctor_get(x_2, 2); +x_11 = lean_ctor_get(x_2, 3); +x_12 = lean_ctor_get(x_2, 4); +x_13 = lean_ctor_get(x_2, 5); +x_14 = lean_ctor_get(x_2, 6); +x_15 = lean_ctor_get(x_2, 7); +x_16 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_2); +x_17 = lean_array_to_list(lean_box(0), x_1); +x_18 = l_List_map___at_Lean_Elab_Command_elabInclude___spec__2(x_17); +x_19 = l_List_appendTR___rarg(x_14, x_18); +x_20 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_20, 0, x_8); +lean_ctor_set(x_20, 1, x_9); +lean_ctor_set(x_20, 2, x_10); +lean_ctor_set(x_20, 3, x_11); +lean_ctor_set(x_20, 4, x_12); +lean_ctor_set(x_20, 5, x_13); +lean_ctor_set(x_20, 6, x_19); +lean_ctor_set(x_20, 7, x_15); +lean_ctor_set_uint8(x_20, sizeof(void*)*8, x_16); +return x_20; +} +} +} +static lean_object* _init_l_Lean_Elab_Command_elabInclude___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("include", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabInclude___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__3; +x_4 = l_Lean_Elab_Command_elabInclude___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = l_Lean_Elab_Command_elabInclude___closed__2; +lean_inc(x_1); +x_6 = l_Lean_Syntax_isOfKind(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); +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; uint8_t x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l_Lean_Syntax_getArgs(x_9); +lean_dec(x_9); +x_11 = l_Lean_Elab_Command_getScope___rarg(x_3, x_4); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 5); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_array_get_size(x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_lt(x_16, x_15); +x_18 = lean_array_size(x_10); +x_19 = 0; +if (x_17 == 0) +{ +lean_object* x_31; +lean_dec(x_15); +lean_dec(x_14); +x_31 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__5___closed__1; +x_20 = x_31; +goto block_30; +} +else +{ +uint8_t x_32; +x_32 = lean_nat_dec_le(x_15, x_15); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_15); +lean_dec(x_14); +x_33 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__5___closed__1; +x_20 = x_33; +goto block_30; +} +else +{ +size_t x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_35 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__5___closed__1; +x_36 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__2(x_14, x_19, x_34, x_35); +lean_dec(x_14); +x_20 = x_36; +goto block_30; +} +} +block_30: +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +lean_inc(x_2); +x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(x_20, x_10, x_18, x_19, x_21, x_2, x_3, x_13); +lean_dec(x_20); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabInclude___lambda__1), 2, 1); +lean_closure_set(x_24, 0, x_10); +x_25 = l_Lean_Elab_Command_modifyScope(x_24, x_2, x_3, x_23); +lean_dec(x_2); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_10); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) +{ +return x_22; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_22, 0); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_22); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Elab_Command_elabInclude(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("elabInclude", 11, 11); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__1; +x_2 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__6; +x_3 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__3; +x_4 = l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabInclude___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabInclude__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__9; +x_3 = l_Lean_Elab_Command_elabInclude___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3; +x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); +return x_6; +} +} static lean_object* _init_l_Lean_logWarning___at_Lean_Elab_Command_elabExit___spec__1___closed__1() { _start: { @@ -36869,6 +37322,27 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabAddDeclDoc_declRange__ if (builtin) {res = l___regBuiltin_Lean_Elab_Command_elabAddDeclDoc_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__4); +l_Lean_Elab_Command_elabInclude___closed__1 = _init_l_Lean_Elab_Command_elabInclude___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_elabInclude___closed__1); +l_Lean_Elab_Command_elabInclude___closed__2 = _init_l_Lean_Elab_Command_elabInclude___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_elabInclude___closed__2); +l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__1); +l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__2); +l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3 = _init_l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabInclude__1___closed__3); +if (builtin) {res = l___regBuiltin_Lean_Elab_Command_elabInclude__1(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); }l_Lean_logWarning___at_Lean_Elab_Command_elabExit___spec__1___closed__1 = _init_l_Lean_logWarning___at_Lean_Elab_Command_elabExit___spec__1___closed__1(); lean_mark_persistent(l_Lean_logWarning___at_Lean_Elab_Command_elabExit___spec__1___closed__1); l_Lean_Elab_Command_elabExit___rarg___closed__1 = _init_l_Lean_Elab_Command_elabExit___rarg___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index 4000cbd66afb..2a73201ed5a5 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -14,7 +14,6 @@ extern "C" { #endif static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_cmdPos___default; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(lean_object*, size_t, size_t, lean_object*); @@ -25,7 +24,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__1(l LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabCommand___spec__22(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607_(lean_object*); lean_object* l_List_iota(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__3; @@ -40,9 +38,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Comm LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__5; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkMessageAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedScope___closed__2; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -66,21 +64,22 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_toArray___at___private_Lean_Elab_Command static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__19; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15; static lean_object* l_Lean_Elab_isIncrementalElab___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__4___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabCommand___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDeclId___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13; LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Language_instInhabitedDynamicSnapshot; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__1; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__6; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___closed__2; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__27; lean_object* lean_private_to_user_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -90,6 +89,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermE LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_liftCommandElabM___rarg___closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(lean_object*, lean_object*); @@ -97,24 +97,23 @@ static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__3; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_opts___default; LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField___at_Lean_Elab_Command_expandDeclId___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__1; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1; static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__1; -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__3; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__9; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4; lean_object* l_Lean_Language_DynamicSnapshot_toTyped_x3f___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__7; lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -131,17 +130,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_tryCatch___rarg(lean_object*, lean_ uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13; uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_traceState___default; static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__2; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_ref___default; -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__23; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandDeclId___spec__17(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___closed__1; @@ -150,10 +148,12 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Comma LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_elabCommand___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); @@ -176,6 +176,7 @@ lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*) static lean_object* l_Lean_Elab_Command_State_scopes___default___closed__1; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13(lean_object*); lean_object* l_String_removeLeadingSpaces(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_toArray___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__16(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -190,27 +191,26 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_messages___default; lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_573_(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3; static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__16; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__6___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_currMacroScope___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__17; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); @@ -227,7 +227,6 @@ lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1; @@ -246,7 +245,6 @@ lean_object* l_Lean_Elab_InfoTree_format(lean_object*, lean_object*, lean_object lean_object* lean_environment_find(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__25; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__12___rarg___boxed(lean_object*, lean_object*); @@ -255,6 +253,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Command_Context_suppressElabErrors___default; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Command___hyg_443_; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -266,9 +265,9 @@ static lean_object* l_Lean_Elab_Command_elabCommand___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; extern lean_object* l_Lean_LocalContext_empty; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -283,6 +282,7 @@ extern lean_object* l_Lean_trace_profiler_useHeartbeats; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -298,6 +298,7 @@ static lean_object* l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM__ LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__20___closed__1; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___closed__1; @@ -308,7 +309,6 @@ uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); @@ -317,12 +317,10 @@ LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___boxed(lean_object*, le lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__14; LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__8; static lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); @@ -331,10 +329,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2___closed__1; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -354,10 +351,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabComma static lean_object* l_Lean_withSetOptionIn___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__4; static lean_object* l_Lean_Elab_Command_elabCommand___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -377,6 +372,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__1___b LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_instToSnapshotTreeMacroExpandedSnapshot___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); @@ -384,7 +380,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_ela LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__17___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_liftCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__16; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM; @@ -401,7 +396,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTerm LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__3(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__28; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1___boxed(lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_commandElabAttribute; @@ -412,11 +407,11 @@ static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8 lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); @@ -424,8 +419,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabComma LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__23; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_nextMacroScope___default; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -433,7 +430,6 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Comman static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576_(lean_object*); uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,7 +449,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCom static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM(lean_object*); -static lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_toList(lean_object*); @@ -465,6 +460,7 @@ static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErro LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4; lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -491,21 +487,19 @@ LEAN_EXPORT lean_object* l_Lean_liftCommandElabM(lean_object*); static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_MetaM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_prefixJoin___at_Lean_Elab_ContextInfo_ppGoals___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_Elab_Command_expandDeclId___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__25(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__26; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__2; -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withScope(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); @@ -519,7 +513,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_scopes___default; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3; static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__3; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__4; lean_object* l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8___boxed(lean_object*, lean_object*, lean_object*); @@ -537,31 +530,31 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_l LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4; lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_withSetOptionIn___closed__4; static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__1; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_ngen___default; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2; extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_Elab_Command_elabCommand___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17; lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___closed__1; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__10; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__1; static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__7; @@ -569,6 +562,7 @@ static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instA LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__19; extern lean_object* l_Lean_Elab_builtinIncrementalElabs; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__7; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -587,13 +581,13 @@ static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErro static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__12___rarg(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__15; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__7(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(size_t, size_t, lean_object*); @@ -604,20 +598,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds___lambda__1(l static lean_object* l_Lean_Elab_Command_State_messages___default___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1; extern lean_object* l_Lean_warningAsError; extern lean_object* l_Lean_Elab_pp_macroStack; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Command___hyg_430_; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__6; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2; static lean_object* l_Lean_Elab_Command_Scope_varDecls___default___closed__1; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__1; extern lean_object* l_Lean_trace_profiler_threshold; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__12; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_liftCommandElabM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -625,23 +616,29 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_levelNames___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM; static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___lambda__2___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__4; static lean_object* l_Lean_liftCommandElabM___rarg___closed__2; static lean_object* l_Lean_Elab_Command_mkState___closed__1; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instToSnapshotTreeMacroExpandedSnapshot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_modifyScope___closed__3; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1; static lean_object* l_Lean_Elab_Command_modifyScope___closed__4; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_currNamespace___default; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__2; extern lean_object* l_Lean_diagnostics; static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__1; static lean_object* l_Lean_Elab_Command_runLinters___lambda__1___closed__1; static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__2; @@ -667,9 +664,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_withLoggingExceptions(lean_object*, static lean_object* l_Lean_Elab_Command_elabCommand___closed__7; lean_object* lean_task_get_own(lean_object*); extern lean_object* l_Std_Format_defWidth; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_expandDeclId___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666_(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_messages___default___closed__3; lean_object* l_Lean_Elab_mkElabAttribute___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -683,12 +683,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__ static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_addLinter(lean_object*, lean_object*); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__3; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5; extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -699,7 +700,6 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCo LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___boxed(lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__13; lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__5; @@ -713,21 +713,22 @@ static lean_object* l_Lean_Elab_Command_modifyScope___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__25; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__26; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_interruptExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__8; static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -739,10 +740,12 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_instToSna LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_elabCommand___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___closed__3; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_586_(lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__2; static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes(lean_object*); @@ -756,18 +759,15 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_openDecls___default; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__24; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; static lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__5; static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1; lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope(lean_object*); @@ -777,11 +777,12 @@ lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16; lean_object* l_Lean_Elab_Term_TermElabM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620_(lean_object*); +static lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkMessageAux(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -799,7 +800,6 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Elab_Com LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__18; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__1; static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__2; @@ -808,6 +808,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Comm static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__3; LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__4; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -816,7 +817,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters_ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1___boxed(lean_object*, lean_object*); @@ -834,6 +834,7 @@ static lean_object* l_Lean_Elab_Command_elabCommand___closed__6; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_modifyScope___spec__1(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_includedVars___default; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__2; @@ -854,7 +855,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCom uint8_t l_Lean_isStructure(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__7___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2; static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -862,8 +862,10 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Core_resetMessageLog(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg___closed__1; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__20; lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__2; lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -879,6 +881,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCom LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__5; lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Command_tryCatch(lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -887,7 +890,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15; lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); extern lean_object* l_Lean_protectedExt; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -898,18 +900,17 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCom lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__3; -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__22; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__7; extern lean_object* l_Lean_firstFrontendMacroScope; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedScope; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM; @@ -920,19 +921,20 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandD LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varUIds___default; uint8_t l_Lean_Syntax_hasMissing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__6; LEAN_EXPORT uint8_t l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1(lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__7; LEAN_EXPORT uint8_t l_Lean_Elab_Command_Scope_isNoncomputable___default; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__5; lean_object* l_Lean_Language_SnapshotTask_map___rarg(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_MessageLog_append(lean_object*, lean_object*); @@ -940,8 +942,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_ioEr LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2; size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope___rarg___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler; @@ -951,7 +953,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandD static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_currRecDepth___default; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__15; lean_object* lean_io_error_to_string(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -962,7 +963,6 @@ static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters__ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2; lean_object* l_Std_Format_prefixJoin___at_Lean_Language_SnapshotTree_format_go___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__2; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -972,6 +972,8 @@ lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg___closed__1; static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_liftCommandElabM___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_List_map___at_Lean_mkConstWithLevelParams___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___boxed(lean_object*); @@ -979,7 +981,9 @@ lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler_output; static lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637_(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__3; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -1016,17 +1020,20 @@ static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__2 LEAN_EXPORT lean_object* l_Lean_Elab_Command_withoutCommandIncrementality(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14; LEAN_EXPORT uint8_t l_Lean_Elab_Command_withoutCommandIncrementality___rarg___lambda__1(uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__4; static lean_object* l_Lean_Elab_Command_modifyScope___closed__1; static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__1; uint8_t l_Lean_Exception_isRuntime(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_SnapshotTree_format_go___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__11; LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1036,10 +1043,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_infoState___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313_(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__1; lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__27; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1058,36 +1063,33 @@ static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab lean_object* lean_dbg_trace(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__1; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_106_(uint8_t, uint8_t); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__10; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_format___at_Lean_Elab_Command_elabCommandTopLevel___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore(lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__24; static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_443____closed__5; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); double lean_float_sub(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Elab_Command___hyg_430____closed__20; LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__3; @@ -1096,7 +1098,6 @@ uint8_t l_Array_isEmpty___rarg(lean_object*); lean_object* l_Lean_Message_toString(lean_object*, uint8_t, lean_object*); extern lean_object* l_Lean_Language_instTypeNameSnapshotLeaf; uint8_t l_Lean_MessageLog_hasErrors(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9; static lean_object* _init_l_Lean_Elab_Command_Scope_opts___default() { _start: { @@ -1146,6 +1147,14 @@ x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; return x_1; } } +static lean_object* _init_l_Lean_Elab_Command_Scope_includedVars___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} static lean_object* _init_l_Lean_Elab_Command_Scope_varUIds___default() { _start: { @@ -1179,15 +1188,16 @@ x_2 = l_Lean_Elab_Command_instInhabitedScope___closed__1; x_3 = lean_box(0); x_4 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; x_5 = 0; -x_6 = lean_alloc_ctor(0, 7, 1); +x_6 = lean_alloc_ctor(0, 8, 1); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_1); lean_ctor_set(x_6, 2, x_3); lean_ctor_set(x_6, 3, x_1); lean_ctor_set(x_6, 4, x_1); lean_ctor_set(x_6, 5, x_4); -lean_ctor_set(x_6, 6, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*7, x_5); +lean_ctor_set(x_6, 6, x_1); +lean_ctor_set(x_6, 7, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*8, x_5); return x_6; } } @@ -1429,7 +1439,7 @@ x_1 = 0; return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__1() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__1() { _start: { lean_object* x_1; @@ -1437,7 +1447,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__2() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__2() { _start: { lean_object* x_1; @@ -1445,7 +1455,7 @@ x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__3() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__3() { _start: { lean_object* x_1; @@ -1453,7 +1463,7 @@ x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__4() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__4() { _start: { lean_object* x_1; @@ -1461,19 +1471,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__5() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__3; -x_4 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__4; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__6() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__6() { _start: { lean_object* x_1; @@ -1481,19 +1491,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__7() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__3; -x_4 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__8() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__8() { _start: { lean_object* x_1; @@ -1501,17 +1511,17 @@ x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__9() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__8; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__10() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__10() { _start: { lean_object* x_1; @@ -1519,41 +1529,41 @@ x_1 = lean_mk_string_unchecked("exact", 5, 5); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__11() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__3; -x_4 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__10; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__10; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__12() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__10; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__10; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__13() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__12; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__14() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__14() { _start: { lean_object* x_1; @@ -1561,7 +1571,7 @@ x_1 = lean_mk_string_unchecked("Term", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__15() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__15() { _start: { lean_object* x_1; @@ -1569,19 +1579,19 @@ x_1 = lean_mk_string_unchecked("declName", 8, 8); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__16() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__14; -x_4 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__15; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__14; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__17() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__17() { _start: { lean_object* x_1; @@ -1589,35 +1599,35 @@ x_1 = lean_mk_string_unchecked("decl_name%", 10, 10); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__18() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__17; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__17; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__19() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__18; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__18; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__20() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__16; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__19; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__16; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__19; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1625,23 +1635,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__21() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__13; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__20; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__13; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__22() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__11; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__21; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__11; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__21; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1649,23 +1659,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__23() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__22; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__22; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__24() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__9; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__23; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__9; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__23; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1673,23 +1683,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__25() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__24; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__24; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__26() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__7; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__25; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__7; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__25; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1697,23 +1707,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__27() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__26; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__26; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__28() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__5; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__27; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__5; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__27; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1721,11 +1731,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_430_() { +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_443_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__28; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__28; return x_1; } } @@ -2122,15 +2132,16 @@ x_6 = lean_box(0); x_7 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; x_8 = 0; lean_inc(x_3); -x_9 = lean_alloc_ctor(0, 7, 1); +x_9 = lean_alloc_ctor(0, 8, 1); lean_ctor_set(x_9, 0, x_5); lean_ctor_set(x_9, 1, x_3); lean_ctor_set(x_9, 2, x_6); lean_ctor_set(x_9, 3, x_4); lean_ctor_set(x_9, 4, x_4); lean_ctor_set(x_9, 5, x_7); -lean_ctor_set(x_9, 6, x_7); -lean_ctor_set_uint8(x_9, sizeof(void*)*7, x_8); +lean_ctor_set(x_9, 6, x_4); +lean_ctor_set(x_9, 7, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*8, x_8); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_4); @@ -2153,7 +2164,7 @@ lean_ctor_set(x_17, 7, x_16); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_573_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_586_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -2179,7 +2190,7 @@ return x_7; } } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1() { _start: { lean_object* x_1; @@ -2187,7 +2198,7 @@ x_1 = lean_mk_string_unchecked("Elab", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2() { _start: { lean_object* x_1; @@ -2195,37 +2206,37 @@ x_1 = lean_mk_string_unchecked("lint", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6() { _start: { lean_object* x_1; @@ -2233,17 +2244,17 @@ x_1 = lean_mk_string_unchecked("Command", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8() { _start: { lean_object* x_1; @@ -2251,17 +2262,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10() { _start: { lean_object* x_1; @@ -2269,47 +2280,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15() { _start: { lean_object* x_1; @@ -2317,33 +2328,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16; -x_2 = lean_unsigned_to_nat(607u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16; +x_2 = lean_unsigned_to_nat(620u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -14048,7 +14059,7 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3; +x_12 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3; x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters___lambda__2), 5, 2); lean_closure_set(x_13, 0, x_1); lean_closure_set(x_13, 1, x_12); @@ -14749,9 +14760,9 @@ static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } @@ -14768,9 +14779,9 @@ static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -14797,7 +14808,7 @@ x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_3, x_4, x_5, x_6, x_7, x_1, x_2); return x_8; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1() { _start: { lean_object* x_1; @@ -14805,23 +14816,23 @@ x_1 = lean_mk_string_unchecked("commandElabAttribute", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2; x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(x_2, x_1); return x_3; } @@ -14993,7 +15004,7 @@ static lean_object* _init_l_Lean_Elab_Command_withoutCommandIncrementality___rar { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Command_withoutCommandIncrementality___rarg___closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; x_3 = l_Lean_Elab_Command_withoutCommandIncrementality___rarg___closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -18267,38 +18278,38 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16; -x_2 = lean_unsigned_to_nat(3313u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16; +x_2 = lean_unsigned_to_nat(3326u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1() { _start: { lean_object* x_1; @@ -18306,23 +18317,23 @@ x_1 = lean_mk_string_unchecked("MacroExpandedSnapshot", 21, 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; -x_4 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; +x_4 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_() { +static lean_object* _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2; +x_1 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2; return x_1; } } @@ -18330,7 +18341,7 @@ static lean_object* _init_l_Lean_Elab_Command_instTypeNameMacroExpandedSnapshot( _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_; +x_1 = l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_; return x_1; } } @@ -21107,7 +21118,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__20___closed__1; x_3 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommand___spec__20___closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); @@ -22734,9 +22745,9 @@ static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -23811,9 +23822,9 @@ static lean_object* _init_l_Lean_Elab_Command_elabCommand___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_Elab_Command_elabCommand___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -23859,7 +23870,7 @@ x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = lean_ctor_get(x_1, 2); lean_inc(x_6); -x_7 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__9; +x_7 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__9; x_8 = lean_name_eq(x_5, x_7); if (x_8 == 0) { @@ -23883,7 +23894,7 @@ lean_inc(x_1); x_13 = l_Lean_Syntax_getKind(x_1); x_14 = 1; x_15 = l_Lean_Name_toString(x_13, x_14); -x_16 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1; +x_16 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1; x_17 = lean_box(x_14); x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___boxed), 8, 5); lean_closure_set(x_18, 0, x_16); @@ -23905,7 +23916,7 @@ x_22 = l_Lean_Syntax_getArg(x_1, x_21); x_23 = l_Lean_Syntax_getKind(x_22); x_24 = 1; x_25 = l_Lean_Name_toString(x_23, x_24); -x_26 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1; +x_26 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1; x_27 = lean_box(x_24); x_28 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___boxed), 8, 5); lean_closure_set(x_28, 0, x_26); @@ -24336,7 +24347,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1() { _start: { lean_object* x_1; @@ -24344,38 +24355,38 @@ x_1 = lean_mk_string_unchecked("input", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16; -x_2 = lean_unsigned_to_nat(4576u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16; +x_2 = lean_unsigned_to_nat(4589u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1() { _start: { lean_object* x_1; @@ -24383,17 +24394,17 @@ x_1 = lean_mk_string_unchecked("showPartialSyntaxErrors", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3() { _start: { lean_object* x_1; @@ -24401,13 +24412,13 @@ x_1 = lean_mk_string_unchecked("show elaboration errors from partial syntax tree return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Lean_Elab_Command_instInhabitedScope___closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -24416,30 +24427,30 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1() { _start: { lean_object* x_1; @@ -24447,27 +24458,27 @@ x_1 = lean_mk_string_unchecked("info", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16; -x_2 = lean_unsigned_to_nat(4653u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16; +x_2 = lean_unsigned_to_nat(4666u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4() { _start: { lean_object* x_1; @@ -24475,23 +24486,23 @@ x_1 = lean_mk_string_unchecked("snapshotTree", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -24499,7 +24510,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; +x_7 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); return x_8; } @@ -25046,7 +25057,7 @@ if (lean_is_exclusive(x_5)) { lean_dec_ref(x_5); x_12 = lean_box(0); } -x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2; +x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2; x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_21, x_6, x_7, x_8); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); @@ -25441,7 +25452,7 @@ if (lean_is_exclusive(x_5)) { lean_dec_ref(x_5); x_12 = lean_box(0); } -x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2; +x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2; x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_21, x_6, x_7, x_8); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); @@ -27441,7 +27452,7 @@ return x_137; else { lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_138 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; +x_138 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; x_139 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_138, x_28, x_3, x_132); x_140 = lean_ctor_get(x_139, 0); lean_inc(x_140); @@ -27742,7 +27753,7 @@ return x_64; else { lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; +x_65 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; x_66 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_65, x_28, x_3, x_59); x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); @@ -28074,7 +28085,7 @@ return x_291; else { lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; -x_292 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; +x_292 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; x_293 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_292, x_28, x_3, x_287); x_294 = lean_ctor_get(x_293, 0); lean_inc(x_294); @@ -28371,7 +28382,7 @@ return x_222; else { lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; -x_223 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5; +x_223 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5; x_224 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(x_223, x_28, x_3, x_218); x_225 = lean_ctor_get(x_224, 0); lean_inc(x_225); @@ -29050,9 +29061,9 @@ static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__2( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__14; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__14; x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -29070,9 +29081,9 @@ static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__14; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__14; x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -29090,9 +29101,9 @@ static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__6( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__14; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__14; x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -29489,7 +29500,7 @@ x_3 = lean_ctor_get(x_2, 2); x_4 = l_Lean_Elab_Command_instInhabitedScope; x_5 = l_List_head_x21___rarg(x_4, x_3); x_6 = lean_box(0); -x_7 = lean_ctor_get(x_5, 6); +x_7 = lean_ctor_get(x_5, 7); lean_inc(x_7); x_8 = lean_array_get_size(x_7); x_9 = lean_unsigned_to_nat(0u); @@ -29547,7 +29558,7 @@ lean_inc(x_19); lean_dec(x_18); x_20 = lean_box(0); x_21 = lean_ctor_get(x_1, 4); -x_22 = lean_ctor_get_uint8(x_5, sizeof(void*)*7); +x_22 = lean_ctor_get_uint8(x_5, sizeof(void*)*8); lean_dec(x_5); x_23 = lean_ctor_get(x_1, 7); x_24 = 1; @@ -32483,7 +32494,7 @@ x_16 = lean_array_get_size(x_3); x_17 = lean_unsigned_to_nat(0u); lean_inc(x_3); x_18 = l_Array_toSubarray___rarg(x_3, x_17, x_16); -x_19 = lean_ctor_get(x_1, 6); +x_19 = lean_ctor_get(x_1, 7); x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_15); @@ -32870,7 +32881,7 @@ static lean_object* _init_l_Lean_Elab_Command_modifyScope___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_Elab_Command_modifyScope___closed__1; x_2 = l_Lean_Elab_Command_modifyScope___closed__2; -x_3 = lean_unsigned_to_nat(673u); +x_3 = lean_unsigned_to_nat(675u); x_4 = lean_unsigned_to_nat(16u); x_5 = l_Lean_Elab_Command_modifyScope___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -34157,7 +34168,7 @@ return x_2; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; x_6 = lean_ctor_get(x_2, 0); x_7 = lean_ctor_get(x_2, 1); x_8 = lean_ctor_get(x_2, 2); @@ -34165,7 +34176,9 @@ x_9 = lean_ctor_get(x_2, 3); x_10 = lean_ctor_get(x_2, 4); x_11 = lean_ctor_get(x_2, 5); x_12 = lean_ctor_get(x_2, 6); -x_13 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_13 = lean_ctor_get(x_2, 7); +x_14 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -34174,19 +34187,20 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_dec(x_2); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_10); -x_15 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_15, 0, x_6); -lean_ctor_set(x_15, 1, x_7); -lean_ctor_set(x_15, 2, x_8); -lean_ctor_set(x_15, 3, x_9); -lean_ctor_set(x_15, 4, x_14); -lean_ctor_set(x_15, 5, x_11); -lean_ctor_set(x_15, 6, x_12); -lean_ctor_set_uint8(x_15, sizeof(void*)*7, x_13); -return x_15; +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_10); +x_16 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_7); +lean_ctor_set(x_16, 2, x_8); +lean_ctor_set(x_16, 3, x_9); +lean_ctor_set(x_16, 4, x_15); +lean_ctor_set(x_16, 5, x_11); +lean_ctor_set(x_16, 6, x_12); +lean_ctor_set(x_16, 7, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*8, x_14); +return x_16; } } } @@ -39145,15 +39159,16 @@ x_22 = l_Lean_Elab_Command_instInhabitedScope___closed__1; x_23 = lean_box(0); x_24 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; lean_inc(x_7); -x_25 = lean_alloc_ctor(0, 7, 1); +x_25 = lean_alloc_ctor(0, 8, 1); lean_ctor_set(x_25, 0, x_22); lean_ctor_set(x_25, 1, x_7); lean_ctor_set(x_25, 2, x_23); lean_ctor_set(x_25, 3, x_16); lean_ctor_set(x_25, 4, x_16); lean_ctor_set(x_25, 5, x_24); -lean_ctor_set(x_25, 6, x_24); -lean_ctor_set_uint8(x_25, sizeof(void*)*7, x_20); +lean_ctor_set(x_25, 6, x_16); +lean_ctor_set(x_25, 7, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*8, x_20); lean_ctor_set_tag(x_11, 1); lean_ctor_set(x_11, 1, x_16); lean_ctor_set(x_11, 0, x_25); @@ -39640,15 +39655,16 @@ x_150 = l_Lean_Elab_Command_instInhabitedScope___closed__1; x_151 = lean_box(0); x_152 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; lean_inc(x_7); -x_153 = lean_alloc_ctor(0, 7, 1); +x_153 = lean_alloc_ctor(0, 8, 1); lean_ctor_set(x_153, 0, x_150); lean_ctor_set(x_153, 1, x_7); lean_ctor_set(x_153, 2, x_151); lean_ctor_set(x_153, 3, x_144); lean_ctor_set(x_153, 4, x_144); lean_ctor_set(x_153, 5, x_152); -lean_ctor_set(x_153, 6, x_152); -lean_ctor_set_uint8(x_153, sizeof(void*)*7, x_148); +lean_ctor_set(x_153, 6, x_144); +lean_ctor_set(x_153, 7, x_152); +lean_ctor_set_uint8(x_153, sizeof(void*)*8, x_148); x_154 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_154, 0, x_153); lean_ctor_set(x_154, 1, x_144); @@ -40851,14 +40867,16 @@ return x_2; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 2); x_7 = lean_ctor_get(x_2, 3); x_8 = lean_ctor_get(x_2, 4); x_9 = lean_ctor_get(x_2, 5); x_10 = lean_ctor_get(x_2, 6); -x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_11 = lean_ctor_get(x_2, 7); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -40866,16 +40884,17 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_1); -lean_ctor_set(x_12, 2, x_6); -lean_ctor_set(x_12, 3, x_7); -lean_ctor_set(x_12, 4, x_8); -lean_ctor_set(x_12, 5, x_9); -lean_ctor_set(x_12, 6, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*7, x_11); -return x_12; +x_13 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_1); +lean_ctor_set(x_13, 2, x_6); +lean_ctor_set(x_13, 3, x_7); +lean_ctor_set(x_13, 4, x_8); +lean_ctor_set(x_13, 5, x_9); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*8, x_12); +return x_13; } } } @@ -40891,9 +40910,9 @@ static lean_object* _init_l_Lean_withSetOptionIn___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_withSetOptionIn___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -40911,9 +40930,9 @@ static lean_object* _init_l_Lean_withSetOptionIn___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__1; -x_2 = l___auto____x40_Lean_Elab_Command___hyg_430____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_443____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6; x_4 = l_Lean_withSetOptionIn___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -41096,6 +41115,8 @@ l_Lean_Elab_Command_Scope_varDecls___default___closed__1 = _init_l_Lean_Elab_Com lean_mark_persistent(l_Lean_Elab_Command_Scope_varDecls___default___closed__1); l_Lean_Elab_Command_Scope_varDecls___default = _init_l_Lean_Elab_Command_Scope_varDecls___default(); lean_mark_persistent(l_Lean_Elab_Command_Scope_varDecls___default); +l_Lean_Elab_Command_Scope_includedVars___default = _init_l_Lean_Elab_Command_Scope_includedVars___default(); +lean_mark_persistent(l_Lean_Elab_Command_Scope_includedVars___default); l_Lean_Elab_Command_Scope_varUIds___default = _init_l_Lean_Elab_Command_Scope_varUIds___default(); lean_mark_persistent(l_Lean_Elab_Command_Scope_varUIds___default); l_Lean_Elab_Command_Scope_isNoncomputable___default = _init_l_Lean_Elab_Command_Scope_isNoncomputable___default(); @@ -41152,64 +41173,64 @@ lean_mark_persistent(l_Lean_Elab_Command_Context_currMacroScope___default); l_Lean_Elab_Command_Context_ref___default = _init_l_Lean_Elab_Command_Context_ref___default(); lean_mark_persistent(l_Lean_Elab_Command_Context_ref___default); l_Lean_Elab_Command_Context_suppressElabErrors___default = _init_l_Lean_Elab_Command_Context_suppressElabErrors___default(); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__1 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__1); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__2 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__2); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__3 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__3); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__4 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__4); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__5 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__5); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__6 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__6); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__7 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__7); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__8 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__8); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__9 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__9); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__10 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__10); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__11 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__11); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__12 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__12); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__13 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__13); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__14 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__14); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__15 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__15); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__16 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__16); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__17 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__17); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__18 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__18); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__19 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__19); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__20 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__20); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__21 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__21); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__22 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__22); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__23 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__23); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__24 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__24); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__25 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__25); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__26 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__26); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__27 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__27); -l___auto____x40_Lean_Elab_Command___hyg_430____closed__28 = _init_l___auto____x40_Lean_Elab_Command___hyg_430____closed__28(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430____closed__28); -l___auto____x40_Lean_Elab_Command___hyg_430_ = _init_l___auto____x40_Lean_Elab_Command___hyg_430_(); -lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_430_); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__1 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__1); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__2 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__2); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__3 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__3); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__4 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__4); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__5 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__5); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__6 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__6); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__7 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__7); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__8 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__8); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__9 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__9); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__10 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__10); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__11 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__11); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__12 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__12); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__13 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__13); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__14 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__14); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__15 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__15); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__16 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__16); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__17 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__17); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__18 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__18); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__19 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__19); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__20 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__20); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__21 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__21); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__22 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__22); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__23 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__23); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__24 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__24); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__25 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__25); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__26 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__26); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__27 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__27); +l___auto____x40_Lean_Elab_Command___hyg_443____closed__28 = _init_l___auto____x40_Lean_Elab_Command___hyg_443____closed__28(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443____closed__28); +l___auto____x40_Lean_Elab_Command___hyg_443_ = _init_l___auto____x40_Lean_Elab_Command___hyg_443_(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_443_); l_Lean_Elab_Command_instMonadCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instMonadCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instMonadCommandElabM___closed__1); l_Lean_Elab_Command_instMonadCommandElabM___closed__2 = _init_l_Lean_Elab_Command_instMonadCommandElabM___closed__2(); @@ -41236,46 +41257,46 @@ l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM = _init_l_Lean_Elab_C lean_mark_persistent(l_Lean_Elab_Command_instMonadExceptOfExceptionCommandElabM); l_Lean_Elab_Command_mkState___closed__1 = _init_l_Lean_Elab_Command_mkState___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkState___closed__1); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_573_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_586_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_lintersRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_lintersRef); lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__14); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__15); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__16); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607____closed__17); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_607_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__15); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__16); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620____closed__17); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_620_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Elab_Command_addLinter___closed__1 = _init_l_Lean_Elab_Command_addLinter___closed__1(); @@ -41455,11 +41476,11 @@ l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7); l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8(); lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624____closed__2); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2624_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637____closed__2); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2637_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_commandElabAttribute = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_commandElabAttribute); @@ -41504,19 +41525,19 @@ l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4 = _init_l_Lean_Ela lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4); l_Lean_Elab_Command_instMonadRecDepthCommandElabM = _init_l_Lean_Elab_Command_instMonadRecDepthCommandElabM(); lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313____closed__2); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3313_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326____closed__2); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3326_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1 = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__1); -l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2 = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414____closed__2); -l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_ = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_(); -lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3414_); +}l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1 = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__1); +l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2 = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427____closed__2); +l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_ = _init_l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_(); +lean_mark_persistent(l_Lean_Elab_Command_instImpl____x40_Lean_Elab_Command___hyg_3427_); l_Lean_Elab_Command_instTypeNameMacroExpandedSnapshot = _init_l_Lean_Elab_Command_instTypeNameMacroExpandedSnapshot(); lean_mark_persistent(l_Lean_Elab_Command_instTypeNameMacroExpandedSnapshot); l_Array_mapMUnsafe_map___at_Lean_Elab_Command_instToSnapshotTreeMacroExpandedSnapshot___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_instToSnapshotTreeMacroExpandedSnapshot___spec__1___closed__1(); @@ -41571,41 +41592,41 @@ l_Lean_Elab_Command_elabCommand___closed__7 = _init_l_Lean_Elab_Command_elabComm lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__7); l_Lean_Elab_Command_elabCommand___boxed__const__1 = _init_l_Lean_Elab_Command_elabCommand___boxed__const__1(); lean_mark_persistent(l_Lean_Elab_Command_elabCommand___boxed__const__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576____closed__3); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4576_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589____closed__3); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4589_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619____closed__5); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4619_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632____closed__5); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4632_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_showPartialSyntaxErrors = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_showPartialSyntaxErrors); lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653____closed__5); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4653_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666____closed__5); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4666_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Frontend.c b/stage0/stdlib/Lean/Elab/Frontend.c index 3043100a2e91..2e1823b7b7dd 100644 --- a/stage0/stdlib/Lean/Elab/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Frontend.c @@ -14,7 +14,9 @@ extern "C" { #endif lean_object* lean_profileit(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommands___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setMessages___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Firefox_Profile_export(lean_object*, double, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_compress(lean_object*); @@ -29,9 +31,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___boxed(lean_object lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); double lean_float_div(double, double); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Elab_process___closed__1; +lean_object* l_String_toNat_x3f(lean_object*); lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3; static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___rarg(lean_object*, lean_object*); @@ -40,16 +45,20 @@ uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setCommandState___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommands(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___closed__1; lean_object* l_Lean_Elab_processHeader(lean_object*, lean_object*, lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos___boxed(lean_object*); extern lean_object* l_Lean_maxRecDepth; +lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +uint8_t lean_string_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Language_SnapshotTree_getAll(lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM(lean_object*); static double l_Lean_Elab_runFrontend___closed__1; static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__6; @@ -65,37 +74,51 @@ lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_Lean_processCommands(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___closed__3; +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2(size_t, size_t, lean_object*); lean_object* l_Lean_Language_reportMessages(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; static lean_object* l_Lean_Elab_process___closed__2; lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setParserState___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_mkInputContext(lean_object*, lean_object*, uint8_t); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__7; +lean_object* l_Lean_getOptionDecls(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5(lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_State_commands___default___closed__1; lean_object* l_Array_toPArray_x27___rarg(lean_object*); +lean_object* l_Lean_Name_getRoot(lean_object*); lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_set_main_module(lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2; static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Util_Profiler_0__Lean_Firefox_toJsonProfile____x40_Lean_Util_Profiler___hyg_3604_(lean_object*); lean_object* l_Lean_Server_ModuleRefs_toLspModuleRefs(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__1; +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___closed__2; LEAN_EXPORT lean_object* lean_run_frontend(lean_object*, lean_object*, lean_object*, lean_object*, uint32_t, lean_object*, uint8_t, lean_object*); lean_object* l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setMessages(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_KVMap_erase(lean_object*, lean_object*); extern lean_object* l_Lean_NameSet_empty; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getCommandState(lean_object*); @@ -103,13 +126,16 @@ static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getInputContext(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommands(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go(lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommandsIncrementally_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_State_commands___default; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__4(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; @@ -124,6 +150,7 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); lean_object* l_Lean_Language_SnapshotTask_get___rarg(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___rarg___boxed(lean_object*, lean_object*); lean_object* l_IO_FS_writeFile(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__5; @@ -135,6 +162,7 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setParserState(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_processCommand___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___boxed(lean_object*); @@ -143,6 +171,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_reparseOptions(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos(lean_object*); static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__2; @@ -3256,6 +3285,804 @@ return x_49; } } } +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, unknown configuration option '", 52, 52); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'\n \nIf the option is defined in this library, use '-D", 62, 62); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' to set it conditionally", 25, 25); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, invalid configuration option '", 52, 52); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' value, it must be true/false", 30, 30); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(1, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(1, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, configuration option '", 44, 44); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' cannot be set in the command line, use set_option command", 59, 59); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' value, it must be a natural number", 36, 36); +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_box(0); +x_10 = l_Lean_Name_replacePrefix(x_1, x_2, x_9); +x_11 = l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(x_3, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_5); +if (x_4 == 0) +{ +uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_6); +x_12 = 1; +lean_inc(x_10); +x_13 = l_Lean_Name_toString(x_10, x_12); +x_14 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_Name_append(x_2, x_10); +x_19 = l_Lean_Name_toString(x_18, x_12); +x_20 = lean_string_append(x_17, x_19); +lean_dec(x_19); +x_21 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_8); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_10); +lean_dec(x_2); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_6); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_8); +return x_26; +} +} +else +{ +uint8_t x_27; +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_11); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_11, 0); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +switch (lean_obj_tag(x_29)) { +case 0: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +lean_ctor_set(x_29, 0, x_5); +x_32 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +lean_ctor_set(x_11, 0, x_32); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_11); +lean_ctor_set(x_33, 1, x_8); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_5); +x_35 = l_Lean_KVMap_insertCore(x_6, x_10, x_34); +lean_ctor_set(x_11, 0, x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_11); +lean_ctor_set(x_36, 1, x_8); +return x_36; +} +} +case 1: +{ +lean_object* x_37; uint8_t x_38; +lean_dec(x_29); +x_37 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; +x_38 = lean_string_dec_eq(x_5, x_37); +if (x_38 == 0) +{ +lean_object* x_39; uint8_t x_40; +x_39 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; +x_40 = lean_string_dec_eq(x_5, x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_10); +lean_dec(x_6); +x_41 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +x_42 = lean_string_append(x_41, x_5); +lean_dec(x_5); +x_43 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; +x_44 = lean_string_append(x_42, x_43); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_44); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_11); +lean_ctor_set(x_45, 1, x_8); +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_5); +x_46 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; +x_47 = l_Lean_KVMap_insertCore(x_6, x_10, x_46); +lean_ctor_set(x_11, 0, x_47); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_11); +lean_ctor_set(x_48, 1, x_8); +return x_48; +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_5); +x_49 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; +x_50 = l_Lean_KVMap_insertCore(x_6, x_10, x_49); +lean_ctor_set(x_11, 0, x_50); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_11); +lean_ctor_set(x_51, 1, x_8); +return x_51; +} +} +case 3: +{ +uint8_t x_52; +x_52 = !lean_is_exclusive(x_29); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_29, 0); +lean_dec(x_53); +x_54 = l_String_toNat_x3f(x_5); +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_free_object(x_29); +lean_dec(x_10); +lean_dec(x_6); +x_55 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +x_56 = lean_string_append(x_55, x_5); +lean_dec(x_5); +x_57 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; +x_58 = lean_string_append(x_56, x_57); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_58); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_11); +lean_ctor_set(x_59, 1, x_8); +return x_59; +} +else +{ +uint8_t x_60; +lean_free_object(x_11); +lean_dec(x_5); +x_60 = !lean_is_exclusive(x_54); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_54, 0); +lean_ctor_set(x_29, 0, x_61); +x_62 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +lean_ctor_set(x_54, 0, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_54); +lean_ctor_set(x_63, 1, x_8); +return x_63; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_54, 0); +lean_inc(x_64); +lean_dec(x_54); +lean_ctor_set(x_29, 0, x_64); +x_65 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_8); +return x_67; +} +} +} +else +{ +lean_object* x_68; +lean_dec(x_29); +x_68 = l_String_toNat_x3f(x_5); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_10); +lean_dec(x_6); +x_69 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +x_70 = lean_string_append(x_69, x_5); +lean_dec(x_5); +x_71 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; +x_72 = lean_string_append(x_70, x_71); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_72); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_11); +lean_ctor_set(x_73, 1, x_8); +return x_73; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_free_object(x_11); +lean_dec(x_5); +x_74 = lean_ctor_get(x_68, 0); +lean_inc(x_74); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + x_75 = x_68; +} else { + lean_dec_ref(x_68); + x_75 = lean_box(0); +} +x_76 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_76, 0, x_74); +x_77 = l_Lean_KVMap_insertCore(x_6, x_10, x_76); +if (lean_is_scalar(x_75)) { + x_78 = lean_alloc_ctor(1, 1, 0); +} else { + x_78 = x_75; +} +lean_ctor_set(x_78, 0, x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_8); +return x_79; +} +} +} +default: +{ +uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_29); +lean_dec(x_6); +lean_dec(x_5); +x_80 = 1; +x_81 = l_Lean_Name_toString(x_10, x_80); +x_82 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; +x_83 = lean_string_append(x_82, x_81); +lean_dec(x_81); +x_84 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; +x_85 = lean_string_append(x_83, x_84); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_85); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_11); +lean_ctor_set(x_86, 1, x_8); +return x_86; +} +} +} +else +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_11, 0); +lean_inc(x_87); +lean_dec(x_11); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +switch (lean_obj_tag(x_88)) { +case 0: +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + x_89 = x_88; +} else { + lean_dec_ref(x_88); + x_89 = lean_box(0); +} +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(0, 1, 0); +} else { + x_90 = x_89; +} +lean_ctor_set(x_90, 0, x_5); +x_91 = l_Lean_KVMap_insertCore(x_6, x_10, x_90); +x_92 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_92, 0, x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_8); +return x_93; +} +case 1: +{ +lean_object* x_94; uint8_t x_95; +lean_dec(x_88); +x_94 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; +x_95 = lean_string_dec_eq(x_5, x_94); +if (x_95 == 0) +{ +lean_object* x_96; uint8_t x_97; +x_96 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; +x_97 = lean_string_dec_eq(x_5, x_96); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_10); +lean_dec(x_6); +x_98 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +x_99 = lean_string_append(x_98, x_5); +lean_dec(x_5); +x_100 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; +x_101 = lean_string_append(x_99, x_100); +x_102 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_8); +return x_103; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_5); +x_104 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; +x_105 = l_Lean_KVMap_insertCore(x_6, x_10, x_104); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_8); +return x_107; +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_5); +x_108 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; +x_109 = l_Lean_KVMap_insertCore(x_6, x_10, x_108); +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_8); +return x_111; +} +} +case 3: +{ +lean_object* x_112; lean_object* x_113; +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + x_112 = x_88; +} else { + lean_dec_ref(x_88); + x_112 = lean_box(0); +} +x_113 = l_String_toNat_x3f(x_5); +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_dec(x_112); +lean_dec(x_10); +lean_dec(x_6); +x_114 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +x_115 = lean_string_append(x_114, x_5); +lean_dec(x_5); +x_116 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; +x_117 = lean_string_append(x_115, x_116); +x_118 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_118, 0, x_117); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_8); +return x_119; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_5); +x_120 = lean_ctor_get(x_113, 0); +lean_inc(x_120); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + x_121 = x_113; +} else { + lean_dec_ref(x_113); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_112)) { + x_122 = lean_alloc_ctor(3, 1, 0); +} else { + x_122 = x_112; +} +lean_ctor_set(x_122, 0, x_120); +x_123 = l_Lean_KVMap_insertCore(x_6, x_10, x_122); +if (lean_is_scalar(x_121)) { + x_124 = lean_alloc_ctor(1, 1, 0); +} else { + x_124 = x_121; +} +lean_ctor_set(x_124, 0, x_123); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_8); +return x_125; +} +} +default: +{ +uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_dec(x_88); +lean_dec(x_6); +lean_dec(x_5); +x_126 = 1; +x_127 = l_Lean_Name_toString(x_10, x_126); +x_128 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; +x_129 = lean_string_append(x_128, x_127); +lean_dec(x_127); +x_130 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; +x_131 = lean_string_append(x_129, x_130); +x_132 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_132, 0, x_131); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_8); +return x_133; +} +} +} +} +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("weak", 4, 4); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_ctor_get(x_7, 0); +lean_inc(x_10); +lean_dec(x_7); +x_11 = l_Lean_Name_getRoot(x_9); +x_12 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2; +x_13 = lean_name_eq(x_11, x_12); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_3, x_14, x_4); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_2 = x_8; +x_3 = x_18; +x_4 = x_17; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_8); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_inc(x_9); +x_24 = l_Lean_KVMap_erase(x_3, x_9); +x_25 = lean_box(0); +x_26 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_24, x_25, x_4); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_2 = x_8; +x_3 = x_29; +x_4 = x_28; +goto _start; +} +else +{ +uint8_t x_31; +lean_dec(x_8); +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) +{ +return x_26; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_26); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +else +{ +lean_object* x_35; +lean_dec(x_7); +lean_dec(x_6); +x_35 = lean_ctor_get(x_2, 1); +lean_inc(x_35); +lean_dec(x_2); +x_2 = x_35; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_reparseOptions(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Lean_getOptionDecls(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +lean_inc(x_1); +x_6 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(x_4, x_1, x_1, x_5); +lean_dec(x_4); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) +{ +return x_6; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_4); +lean_dec(x_4); +x_10 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -3728,7 +4555,7 @@ x_26 = l_Lean_Elab_processHeader(x_23, x_2, x_25, x_18, x_5, x_17, x_22); lean_dec(x_23); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; double x_36; double x_37; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); @@ -3739,190 +4566,228 @@ lean_inc(x_29); x_30 = lean_ctor_get(x_27, 1); lean_inc(x_30); lean_dec(x_27); +x_31 = l_Lean_Elab_reparseOptions(x_2, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; double x_39; double x_40; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_4); -x_31 = lean_environment_set_main_module(x_29, x_4); -lean_inc(x_2); +x_34 = lean_environment_set_main_module(x_29, x_4); +lean_inc(x_32); lean_inc(x_30); -lean_inc(x_31); -x_32 = l_Lean_Elab_Command_mkState(x_31, x_30, x_2); -x_33 = lean_io_mono_nanos_now(x_28); -x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Float_ofScientific(x_34, x_12, x_13); -lean_dec(x_34); -x_37 = lean_float_div(x_36, x_15); +x_35 = l_Lean_Elab_Command_mkState(x_34, x_30, x_32); +x_36 = lean_io_mono_nanos_now(x_33); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Float_ofScientific(x_37, x_12, x_13); +lean_dec(x_37); +x_40 = lean_float_div(x_39, x_15); if (lean_obj_tag(x_6) == 0) { -lean_object* x_38; lean_object* x_39; -lean_dec(x_31); +lean_object* x_41; lean_object* x_42; +lean_dec(x_34); lean_dec(x_30); -x_38 = lean_box(0); -x_39 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_2, x_7, x_16, x_37, x_4, x_6, x_32, x_38, x_35); -return x_39; +x_41 = lean_box(0); +x_42 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_41, x_38); +return x_42; } else { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_32); -if (x_40 == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_35); +if (x_43 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_41 = lean_ctor_get(x_32, 6); -x_42 = lean_ctor_get(x_32, 4); -lean_dec(x_42); -x_43 = lean_ctor_get(x_32, 3); -lean_dec(x_43); -x_44 = lean_ctor_get(x_32, 1); -lean_dec(x_44); -x_45 = lean_ctor_get(x_32, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_44 = lean_ctor_get(x_35, 6); +x_45 = lean_ctor_get(x_35, 4); lean_dec(x_45); -x_46 = l_Lean_Elab_runFrontend___closed__3; -x_47 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_46); -x_48 = !lean_is_exclusive(x_41); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_ctor_set_uint8(x_41, sizeof(void*)*2, x_17); -x_49 = l_Lean_Elab_runFrontend___closed__2; -lean_ctor_set(x_32, 4, x_47); -lean_ctor_set(x_32, 3, x_49); -lean_ctor_set(x_32, 1, x_30); -lean_ctor_set(x_32, 0, x_31); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_2, x_7, x_16, x_37, x_4, x_6, x_32, x_50, x_35); +x_46 = lean_ctor_get(x_35, 3); +lean_dec(x_46); +x_47 = lean_ctor_get(x_35, 1); +lean_dec(x_47); +x_48 = lean_ctor_get(x_35, 0); +lean_dec(x_48); +x_49 = l_Lean_Elab_runFrontend___closed__3; +x_50 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_32, x_49); +x_51 = !lean_is_exclusive(x_44); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_ctor_set_uint8(x_44, sizeof(void*)*2, x_17); +x_52 = l_Lean_Elab_runFrontend___closed__2; +lean_ctor_set(x_35, 4, x_50); +lean_ctor_set(x_35, 3, x_52); +lean_ctor_set(x_35, 1, x_30); +lean_ctor_set(x_35, 0, x_34); +x_53 = lean_box(0); +x_54 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_53, x_38); lean_dec(x_6); -return x_51; +return x_54; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_41, 0); -x_53 = lean_ctor_get(x_41, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_41); -x_54 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set_uint8(x_54, sizeof(void*)*2, x_17); -x_55 = l_Lean_Elab_runFrontend___closed__2; -lean_ctor_set(x_32, 6, x_54); -lean_ctor_set(x_32, 4, x_47); -lean_ctor_set(x_32, 3, x_55); -lean_ctor_set(x_32, 1, x_30); -lean_ctor_set(x_32, 0, x_31); -x_56 = lean_box(0); -x_57 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_2, x_7, x_16, x_37, x_4, x_6, x_32, x_56, x_35); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_55 = lean_ctor_get(x_44, 0); +x_56 = lean_ctor_get(x_44, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_44); +x_57 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set_uint8(x_57, sizeof(void*)*2, x_17); +x_58 = l_Lean_Elab_runFrontend___closed__2; +lean_ctor_set(x_35, 6, x_57); +lean_ctor_set(x_35, 4, x_50); +lean_ctor_set(x_35, 3, x_58); +lean_ctor_set(x_35, 1, x_30); +lean_ctor_set(x_35, 0, x_34); +x_59 = lean_box(0); +x_60 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_59, x_38); lean_dec(x_6); -return x_57; +return x_60; } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_58 = lean_ctor_get(x_32, 2); -x_59 = lean_ctor_get(x_32, 5); -x_60 = lean_ctor_get(x_32, 6); -x_61 = lean_ctor_get(x_32, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_32); -x_62 = l_Lean_Elab_runFrontend___closed__3; -x_63 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_62); -x_64 = lean_ctor_get(x_60, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_61 = lean_ctor_get(x_35, 2); +x_62 = lean_ctor_get(x_35, 5); +x_63 = lean_ctor_get(x_35, 6); +x_64 = lean_ctor_get(x_35, 7); lean_inc(x_64); -x_65 = lean_ctor_get(x_60, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_66 = x_60; +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_35); +x_65 = l_Lean_Elab_runFrontend___closed__3; +x_66 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_32, x_65); +x_67 = lean_ctor_get(x_63, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_63, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_69 = x_63; } else { - lean_dec_ref(x_60); - x_66 = lean_box(0); + lean_dec_ref(x_63); + x_69 = lean_box(0); } -if (lean_is_scalar(x_66)) { - x_67 = lean_alloc_ctor(0, 2, 1); +if (lean_is_scalar(x_69)) { + x_70 = lean_alloc_ctor(0, 2, 1); } else { - x_67 = x_66; -} -lean_ctor_set(x_67, 0, x_64); -lean_ctor_set(x_67, 1, x_65); -lean_ctor_set_uint8(x_67, sizeof(void*)*2, x_17); -x_68 = l_Lean_Elab_runFrontend___closed__2; -x_69 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_69, 0, x_31); -lean_ctor_set(x_69, 1, x_30); -lean_ctor_set(x_69, 2, x_58); -lean_ctor_set(x_69, 3, x_68); -lean_ctor_set(x_69, 4, x_63); -lean_ctor_set(x_69, 5, x_59); -lean_ctor_set(x_69, 6, x_67); -lean_ctor_set(x_69, 7, x_61); -x_70 = lean_box(0); -x_71 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_2, x_7, x_16, x_37, x_4, x_6, x_69, x_70, x_35); + x_70 = x_69; +} +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +lean_ctor_set_uint8(x_70, sizeof(void*)*2, x_17); +x_71 = l_Lean_Elab_runFrontend___closed__2; +x_72 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_72, 0, x_34); +lean_ctor_set(x_72, 1, x_30); +lean_ctor_set(x_72, 2, x_61); +lean_ctor_set(x_72, 3, x_71); +lean_ctor_set(x_72, 4, x_66); +lean_ctor_set(x_72, 5, x_62); +lean_ctor_set(x_72, 6, x_70); +lean_ctor_set(x_72, 7, x_64); +x_73 = lean_box(0); +x_74 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_72, x_73, x_38); lean_dec(x_6); -return x_71; +return x_74; +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_6); +lean_dec(x_4); +x_75 = !lean_is_exclusive(x_31); +if (x_75 == 0) +{ +return x_31; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_31, 0); +x_77 = lean_ctor_get(x_31, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_31); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } } else { -uint8_t x_72; +uint8_t x_79; lean_dec(x_24); lean_dec(x_18); lean_dec(x_6); lean_dec(x_4); lean_dec(x_2); -x_72 = !lean_is_exclusive(x_26); -if (x_72 == 0) +x_79 = !lean_is_exclusive(x_26); +if (x_79 == 0) { return x_26; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_26, 0); -x_74 = lean_ctor_get(x_26, 1); -lean_inc(x_74); -lean_inc(x_73); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_26, 0); +x_81 = lean_ctor_get(x_26, 1); +lean_inc(x_81); +lean_inc(x_80); lean_dec(x_26); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } } else { -uint8_t x_76; +uint8_t x_83; lean_dec(x_18); lean_dec(x_6); lean_dec(x_4); lean_dec(x_2); -x_76 = !lean_is_exclusive(x_19); -if (x_76 == 0) +x_83 = !lean_is_exclusive(x_19); +if (x_83 == 0) { return x_19; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_19, 0); -x_78 = lean_ctor_get(x_19, 1); -lean_inc(x_78); -lean_inc(x_77); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_19, 0); +x_85 = lean_ctor_get(x_19, 1); +lean_inc(x_85); +lean_inc(x_84); lean_dec(x_19); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } @@ -4024,6 +4889,34 @@ l_Lean_Elab_process___closed__1 = _init_l_Lean_Elab_process___closed__1(); lean_mark_persistent(l_Lean_Elab_process___closed__1); l_Lean_Elab_process___closed__2 = _init_l_Lean_Elab_process___closed__2(); lean_mark_persistent(l_Lean_Elab_process___closed__2); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1); +l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2); l_Lean_Elab_runFrontend___lambda__2___closed__1 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__1); l_Lean_Elab_runFrontend___lambda__2___closed__2 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Import.c b/stage0/stdlib/Lean/Elab/Import.c index 2c7c97f5ee90..ff8f6f6d5985 100644 --- a/stage0/stdlib/Lean/Elab/Import.c +++ b/stage0/stdlib/Lean/Elab/Import.c @@ -42,7 +42,7 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_processHeader___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_headerToImports___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_parseImports___closed__1; -lean_object* l_Lean_findOLean(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_findOLean(lean_object*, lean_object*); static lean_object* l_Lean_Elab_headerToImports___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_parseImports(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -894,80 +894,79 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_dec(x_4); x_8 = lean_array_uget(x_1, x_3); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); lean_dec(x_8); -x_10 = 0; -x_11 = l_Lean_findOLean(x_9, x_10, x_5); -if (lean_obj_tag(x_11) == 0) +x_10 = l_Lean_findOLean(x_9, x_5); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_IO_println___at_Lean_Elab_printImports___spec__1(x_12, x_13); -if (lean_obj_tag(x_14) == 0) +lean_dec(x_10); +x_13 = l_IO_println___at_Lean_Elab_printImports___spec__1(x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_18 = lean_box(0); -x_3 = x_17; -x_4 = x_18; -x_5 = x_15; +lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_17 = lean_box(0); +x_3 = x_16; +x_4 = x_17; +x_5 = x_14; goto _start; } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_13); +if (x_19 == 0) { -return x_14; +return x_13; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); lean_inc(x_21); -lean_dec(x_14); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } else { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_11); -if (x_24 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_10); +if (x_23 == 0) { -return x_11; +return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_10, 0); +x_25 = lean_ctor_get(x_10, 1); lean_inc(x_25); -lean_dec(x_11); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_inc(x_24); +lean_dec(x_10); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index cb14832b71f6..06a16866d3aa 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -34295,7 +34295,7 @@ return x_1; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; x_10 = lean_ctor_get(x_1, 0); x_11 = lean_ctor_get(x_1, 1); x_12 = lean_ctor_get(x_1, 2); @@ -34303,7 +34303,9 @@ x_13 = lean_ctor_get(x_1, 3); x_14 = lean_ctor_get(x_1, 4); x_15 = lean_ctor_get(x_1, 5); x_16 = lean_ctor_get(x_1, 6); -x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +x_17 = lean_ctor_get(x_1, 7); +x_18 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); @@ -34312,22 +34314,23 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_dec(x_1); -x_18 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__2; -x_19 = 0; -x_20 = l_Lean_KVMap_setBool(x_11, x_18, x_19); -x_21 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__4; -x_22 = 1; -x_23 = l_Lean_KVMap_setBool(x_20, x_21, x_22); -x_24 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_24, 0, x_10); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_12); -lean_ctor_set(x_24, 3, x_13); -lean_ctor_set(x_24, 4, x_14); -lean_ctor_set(x_24, 5, x_15); -lean_ctor_set(x_24, 6, x_16); -lean_ctor_set_uint8(x_24, sizeof(void*)*7, x_17); -return x_24; +x_19 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__2; +x_20 = 0; +x_21 = l_Lean_KVMap_setBool(x_11, x_19, x_20); +x_22 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__4; +x_23 = 1; +x_24 = l_Lean_KVMap_setBool(x_21, x_22, x_23); +x_25 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_12); +lean_ctor_set(x_25, 3, x_13); +lean_ctor_set(x_25, 4, x_14); +lean_ctor_set(x_25, 5, x_15); +lean_ctor_set(x_25, 6, x_16); +lean_ctor_set(x_25, 7, x_17); +lean_ctor_set_uint8(x_25, sizeof(void*)*8, x_18); +return x_25; } } } diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 0d3c4325a93c..71c7a071694b 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -22,18 +22,19 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__4___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1; lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_levelMVarToParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getKindForLetRecs___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_iota(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4(lean_object*); @@ -45,6 +46,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeader LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getAllUserLevelNames(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -53,6 +55,7 @@ static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___ lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -64,10 +67,10 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instFVarIdSetInhabited; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___boxed(lean_object**); @@ -87,44 +90,48 @@ static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__6; LEAN_EXPORT lean_object* l_Array_erase___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__2(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_merge(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkBodyTask(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__2; -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__5; extern lean_object* l_Lean_declRangeExt; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__5; static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1; lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9; extern lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__8___boxed(lean_object**); -static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_newLetDecls___default; LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_DynamicSnapshot_toTyped_x3f___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -145,29 +152,31 @@ uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_ob static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__7; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1___boxed(lean_object**); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_processDeriving(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__11; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___closed__1; lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_elabCommand___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4; @@ -185,6 +194,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_ lean_object* l_Lean_LocalDecl_replaceFVarId(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_pushMain___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___lambda__3(lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -200,24 +210,23 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_deprecated_oldSectionVars; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__1; uint8_t l_Array_contains___at___private_Lean_Class_0__Lean_checkOutParam___spec__1(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_newLocalDecls___default; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___lambda__1___boxed(lean_object*, lean_object*); @@ -229,68 +238,73 @@ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__14; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___boxed(lean_object**); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__3___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__11___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); -static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__9; -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546_(lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__6(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_DefKind_isDefOrAbbrevOrOpaque(uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Elab_Command_elabMutualDef___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3(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_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample(lean_object*); -static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___closed__3; lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__15___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__5(lean_object*); static lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__6(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__2; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_merge___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,24 +329,30 @@ extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply(lean_object*, lean_object*); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__2(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5(lean_object*, size_t, size_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__1; lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__3; +lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_getAndEmptyMessageLog___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__4; @@ -343,6 +363,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_level LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); uint8_t l_Lean_Elab_DefKind_isExample(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_elabSyntheticHole___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); @@ -352,6 +373,7 @@ lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -361,7 +383,6 @@ lean_object* l_Lean_LocalDecl_index(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandLetEqnsDecl(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__6___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); @@ -370,30 +391,25 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___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_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__6; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4; static lean_object* l_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2___closed__2; static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9; lean_object* lean_io_promise_resolve(lean_object*, lean_object*, lean_object*); @@ -408,9 +424,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_insta LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint(lean_object*); lean_object* l_Lean_Elab_Term_expandWhereDeclsOpt(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Command_elabMutualDef___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); @@ -422,7 +438,6 @@ static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_ lean_object* l_Lean_Meta_resetZetaDeltaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__1(lean_object*); @@ -438,20 +453,22 @@ lean_object* l_Lean_LocalDecl_setUserName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1; static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1___closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getDeclarationSelectionRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_isApplicable___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_realizeGlobalConstNoOverload(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_MutualClosure_getKindForLetRecs(lean_object*); @@ -466,25 +483,25 @@ lean_object* l_Array_unzip___rarg(lean_object*); lean_object* l_Array_feraseIdx___rarg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__12; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_findSome_x3f___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14___boxed(lean_object**); lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified___boxed(lean_object*); uint8_t l_Lean_RBNode_isSingleton___rarg(lean_object*); uint8_t l_Lean_Elab_DefKind_isTheorem(uint8_t); LEAN_EXPORT uint8_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___closed__2; lean_object* lean_io_add_heartbeats(uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___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* l_Lean_mkHole(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); @@ -496,13 +513,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_El LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_elabMutualDef___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -516,11 +530,13 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__5; +lean_object* l_Lean_CollectFVars_State_addDependencies(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_MutualClosure_insertReplacementForLetRecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_MutualClosure_isApplicable___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Elab_Term_applyAttributesAt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -534,12 +550,10 @@ lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__4; lean_object* l_Lean_HashMap_insert___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2; lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -548,6 +562,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___lambda__1(l lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_indexOfAux___at_Lean_LocalContext_erase___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7(lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -557,25 +572,28 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_El LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__3; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2; uint8_t l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_17_(uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__12___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_universes; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3___boxed(lean_object**); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -586,6 +604,7 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkForHiddenUnivLevel static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__4; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_MutualClosure_main___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -602,23 +621,23 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isMissing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_localDecls___default; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__3; static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; -static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__1___boxed(lean_object**); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__13; @@ -631,39 +650,41 @@ lean_object* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5(lean_object*); lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__8; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6; +uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); uint8_t l_Lean_Elab_DefView_isInstance(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1; +lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_markModified___boxed(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; lean_object* l_Lean_Syntax_getSepArgs(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2; lean_object* l_Lean_Elab_Term_getLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__10___boxed(lean_object**); @@ -679,6 +700,7 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLet uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_erase___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__2(lean_object*, lean_object*); @@ -711,6 +733,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0 static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5; lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_resetModified___rarg(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -718,23 +741,27 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_ lean_object* l_Lean_Expr_appFn_x21(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); extern lean_object* l_Lean_Language_Snapshot_Diagnostics_empty; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader; lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* lean_find_ext_expr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_preprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717_(lean_object*); lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__2; lean_object* l_Lean_Meta_getZetaDeltaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_getUsedFVarsMap___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__8(lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Closure_mkLambda(lean_object*, lean_object*); @@ -751,6 +778,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_isApplicable___boxed(lea LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_getKindForLetRecs___boxed(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_instantiateMVarsProfiling(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__5; @@ -769,6 +797,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabM static lean_object* l_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__12___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___closed__2; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__8; @@ -782,34 +811,42 @@ lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_ob LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__8___lambda__1___boxed(lean_object**); lean_object* l_Lean_Elab_Command_mkDefView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7; extern lean_object* l_Lean_Elab_instTypeNameDefsParsedSnapshot; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___closed__1; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__2; lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_TerminationHints_rememberExtraParams(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12; lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -820,12 +857,14 @@ static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualD static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__4; +uint8_t l_Lean_Name_hasMacroScopes(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addAutoBoundImplicits(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__2___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_processDeriving___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_insertReplacementForMainFns___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap(lean_object*); @@ -846,27 +885,30 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendind LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__5; lean_object* l_Lean_Elab_addPreDefinitions(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___boxed(lean_object**); static lean_object* l_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2___closed__1; static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_CollectMVars_visit(lean_object*, lean_object*); +static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___lambda__1___boxed(lean_object*, lean_object*); uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_433____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___lambda__2___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint___rarg(lean_object*, lean_object*); lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__10; static lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -874,8 +916,10 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0_ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -886,16 +930,17 @@ lean_object* l_Lean_MessageData_ofLevel(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_pushLetRecs(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__5; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__3(lean_object*, size_t, size_t); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__2; @@ -904,7 +949,6 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClos LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__9___boxed(lean_object**); lean_object* l_Lean_indentExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_funBinderTypes; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -916,7 +960,7 @@ LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_push LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__3; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); @@ -938,21 +982,25 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualD LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_merge___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3; static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__4; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_modifyUsedFVars(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Elab_addPreDefinitions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_markModified(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getDocStringText___at_Lean_Elab_Command_elabMutualDef___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -961,14 +1009,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualD static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1; static size_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__13; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_get_x21(lean_object*, lean_object*); @@ -979,9 +1026,12 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosur static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___closed__4; uint8_t l_Lean_Expr_isProp(lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1___boxed(lean_object**); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_isApplicable___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__3; @@ -989,6 +1039,7 @@ static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambd LEAN_EXPORT lean_object* l_Lean_getDocStringText___at_Lean_Elab_Command_elabMutualDef___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); @@ -1002,21 +1053,27 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeader lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__15___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_Lean_Elab_Term_withAuxDecl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasSorry(lean_object*); +static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; lean_object* l_String_intercalate(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__8; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__10; extern lean_object* l_Lean_Elab_Command_instInhabitedScope; @@ -1024,7 +1081,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint___boxed(lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6; LEAN_EXPORT lean_object* l_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__1(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -1038,7 +1094,7 @@ uint8_t l_Lean_Syntax_hasMissing(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__6___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getDocStringText___at_Lean_Elab_Command_elabMutualDef___spec__9___closed__1; lean_object* l_Lean_getDelayedMVarRoot___rarg(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift; @@ -1052,17 +1108,20 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfN LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasFVar(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2(lean_object*, lean_object*, size_t, size_t); +lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_State_usedFVarsMap___default; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11; static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1088,7 +1147,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__3; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__13___boxed(lean_object*, lean_object*, lean_object*); @@ -1104,12 +1163,15 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__9; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__7; lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_map___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_get_size(lean_object*); extern lean_object* l_Lean_pp_letVarTypes; @@ -1121,9 +1183,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_Mutua LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__13(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1139,11 +1204,11 @@ uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_resetModified___boxed(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample___boxed(lean_object*); -static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2; static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__12___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1152,21 +1217,21 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMul static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__9; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___boxed(lean_object**); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3; static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3; uint8_t l_Lean_Exception_isRuntime(lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__9; uint8_t l_Lean_KVMap_contains(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___boxed__const__1; -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_CollectFVars_State_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__1; @@ -1175,18 +1240,20 @@ static lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__4; lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__9; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_fixLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandWhereDecls(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642_(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5; @@ -1203,7 +1270,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutual LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at_Lean_Elab_Command_elabMutualDef___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldTR_loop___at_Lean_Elab_Term_MutualClosure_insertReplacementForMainFns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getLetRecsToLift___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1213,6 +1279,7 @@ lean_object* l_Lean_Expr_collectFVars(lean_object*, lean_object*, lean_object*, LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at_Lean_Elab_Command_elabMutualDef___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__5; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1221,17 +1288,22 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_markModified___rarg(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified(lean_object*); LEAN_EXPORT lean_object* l_Nat_foldTR_loop___at_Lean_Elab_Term_MutualClosure_insertReplacementForMainFns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__12___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasExprMVar(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_instantiateMVarsAtHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1239,7 +1311,7 @@ static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualD LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_MessageLog_hasErrors(lean_object*); lean_object* l_Lean_Elab_Term_withDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1() { _start: { @@ -17984,731 +18056,3792 @@ lean_dec(x_2); return x_10; } } +static lean_object* _init_l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instantiate metavars", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_instantiateMVarsProfiling(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_4, 2); +lean_inc(x_7); +x_8 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); +lean_closure_set(x_8, 0, x_1); +x_9 = l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1; +x_10 = lean_box(0); +x_11 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__6___rarg(x_9, x_7, x_8, x_10, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_7); +return x_11; +} +} +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = l_Lean_mkHashSetImp___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; +x_3 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; +x_8 = lean_st_mk_ref(x_7, x_6); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Expr_collectFVars(x_1, x_9, x_2, x_3, x_4, x_5, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_9, x_12); +lean_dec(x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 2); +lean_inc(x_16); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_ctor_get(x_17, 2); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_3, x_4); +x_8 = l_Lean_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_visit___spec__1(x_2, x_7); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = 1; +return x_9; +} +else +{ +size_t x_10; size_t x_11; +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_4 = x_11; +goto _start; +} +} +else +{ +uint8_t x_13; +x_13 = 0; +return x_13; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = l_Lean_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_visit___spec__1(x_1, x_6); +lean_dec(x_6); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +else +{ +size_t x_9; size_t x_10; +lean_dec(x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; +goto _start; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_eq(x_2, x_3); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_12 = lean_array_uget(x_1, x_2); +lean_inc(x_6); +x_13 = l_Lean_Meta_getFVarLocalDecl(x_12, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_36; uint8_t x_37; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_5, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_LocalDecl_type(x_14); +x_20 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars(x_19, x_6, x_7, x_8, x_9, x_18); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_36 = l_Lean_LocalDecl_binderInfo(x_14); +x_37 = l_Lean_BinderInfo_isInstImplicit(x_36); +if (x_37 == 0) +{ +size_t x_38; size_t x_39; lean_object* x_40; +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_14); +x_38 = 1; +x_39 = lean_usize_add(x_2, x_38); +x_40 = lean_box(0); +x_2 = x_39; +x_4 = x_40; +x_10 = x_22; +goto _start; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_42 = lean_ctor_get(x_17, 1); +lean_inc(x_42); +lean_dec(x_17); +x_43 = lean_array_get_size(x_21); +x_44 = lean_unsigned_to_nat(0u); +x_45 = lean_nat_dec_lt(x_44, x_43); +if (x_45 == 0) +{ +lean_object* x_46; +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_21); +x_46 = lean_box(0); +x_23 = x_46; +goto block_35; +} +else +{ +size_t x_47; size_t x_48; uint8_t x_49; +x_47 = 0; +x_48 = lean_usize_of_nat(x_43); +lean_dec(x_43); +x_49 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2(x_42, x_21, x_47, x_48); +lean_dec(x_21); +lean_dec(x_42); +if (x_49 == 0) +{ +lean_object* x_50; +x_50 = lean_box(0); +x_23 = x_50; +goto block_35; +} +else +{ +size_t x_51; size_t x_52; lean_object* x_53; +lean_dec(x_14); +x_51 = 1; +x_52 = lean_usize_add(x_2, x_51); +x_53 = lean_box(0); +x_2 = x_52; +x_4 = x_53; +x_10 = x_22; +goto _start; +} +} +} +block_35: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; +lean_dec(x_23); +x_24 = lean_st_ref_take(x_5, x_22); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_LocalDecl_fvarId(x_14); +lean_dec(x_14); +x_28 = l_Lean_CollectFVars_State_add(x_25, x_27); +x_29 = lean_st_ref_set(x_5, x_28, x_26); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = 1; +x_32 = lean_usize_add(x_2, x_31); +x_33 = lean_box(0); +x_2 = x_32; +x_4 = x_33; +x_10 = x_30; +goto _start; +} +} +else +{ +uint8_t x_55; +lean_dec(x_6); +x_55 = !lean_is_exclusive(x_13); +if (x_55 == 0) +{ +return x_13; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_13, 0); +x_57 = lean_ctor_get(x_13, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_13); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; +lean_dec(x_6); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_4); +lean_ctor_set(x_59, 1, x_10); +return x_59; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = lean_array_uget(x_2, x_3); +lean_inc(x_7); +x_14 = l_Lean_Meta_getFVarLocalDecl(x_13, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_LocalDecl_userName(x_15); +x_18 = l_List_elem___at_Lean_NameHashSet_insert___spec__2(x_17, x_1); +lean_dec(x_17); +if (x_18 == 0) +{ +size_t x_19; size_t x_20; lean_object* x_21; +lean_dec(x_15); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_21 = lean_box(0); +x_3 = x_20; +x_5 = x_21; +x_11 = x_16; +goto _start; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; +x_23 = lean_st_ref_take(x_6, x_16); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_LocalDecl_fvarId(x_15); +lean_dec(x_15); +x_27 = l_Lean_CollectFVars_State_add(x_24, x_26); +x_28 = lean_st_ref_set(x_6, x_27, x_25); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = 1; +x_31 = lean_usize_add(x_3, x_30); +x_32 = lean_box(0); +x_3 = x_31; +x_5 = x_32; +x_11 = x_29; +goto _start; +} +} +else +{ +uint8_t x_34; +lean_dec(x_7); +x_34 = !lean_is_exclusive(x_14); +if (x_34 == 0) +{ +return x_14; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_14, 0); +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_14); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; +lean_dec(x_7); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_5); +lean_ctor_set(x_38, 1, x_11); +return x_38; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_eq(x_2, x_3); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; +lean_dec(x_4); +x_12 = lean_array_uget(x_1, x_2); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 5); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_Expr_collectFVars(x_14, x_5, x_6, x_7, x_8, x_9, x_10); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = 1; +x_19 = lean_usize_add(x_2, x_18); +x_2 = x_19; +x_4 = x_16; +x_10 = x_17; +goto _start; +} +else +{ +lean_object* x_21; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_4); +lean_ctor_set(x_21, 1, x_10); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_10 = lean_array_get_size(x_3); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_10); +if (x_12 == 0) +{ +lean_dec(x_10); +x_13 = x_9; +goto block_58; +} +else +{ +uint8_t x_59; +x_59 = lean_nat_dec_le(x_10, x_10); +if (x_59 == 0) +{ +lean_dec(x_10); +x_13 = x_9; +goto block_58; +} +else +{ +size_t x_60; size_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = 0; +x_61 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_62 = lean_box(0); +x_63 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5(x_3, x_60, x_61, x_62, x_4, x_5, x_6, x_7, x_8, x_9); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_13 = x_64; +goto block_58; +} +} +block_58: +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_array_get_size(x_1); +x_15 = lean_nat_dec_lt(x_11, x_14); +if (x_15 == 0) +{ +x_16 = x_13; +goto block_47; +} +else +{ +uint8_t x_48; +x_48 = lean_nat_dec_le(x_14, x_14); +if (x_48 == 0) +{ +x_16 = x_13; +goto block_47; +} +else +{ +size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +x_49 = 0; +x_50 = lean_usize_of_nat(x_14); +x_51 = lean_box(0); +lean_inc(x_5); +x_52 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4(x_2, x_1, x_49, x_50, x_51, x_4, x_5, x_6, x_7, x_8, x_13); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_16 = x_53; +goto block_47; +} +else +{ +uint8_t x_54; +lean_dec(x_14); +lean_dec(x_5); +x_54 = !lean_is_exclusive(x_52); +if (x_54 == 0) +{ +return x_52; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_52, 0); +x_56 = lean_ctor_get(x_52, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_52); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +} +block_47: +{ +lean_object* x_17; 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_st_ref_get(x_4, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_inc(x_5); +x_20 = l_Lean_CollectFVars_State_addDependencies(x_18, x_5, x_6, x_7, x_8, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_st_ref_set(x_4, x_21, x_22); +if (x_15 == 0) +{ +uint8_t x_24; +lean_dec(x_14); +lean_dec(x_5); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_23, 0, x_26); +return x_23; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_ctor_get(x_23, 1); +x_32 = lean_ctor_get(x_23, 0); +lean_dec(x_32); +x_33 = lean_nat_dec_le(x_14, x_14); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_14); +lean_dec(x_5); +x_34 = lean_box(0); +lean_ctor_set(x_23, 0, x_34); +return x_23; +} +else +{ +size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +lean_free_object(x_23); +x_35 = 0; +x_36 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_37 = lean_box(0); +x_38 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3(x_1, x_35, x_36, x_37, x_4, x_5, x_6, x_7, x_8, x_31); +return x_38; +} +} +else +{ +lean_object* x_39; uint8_t x_40; +x_39 = lean_ctor_get(x_23, 1); +lean_inc(x_39); +lean_dec(x_23); +x_40 = lean_nat_dec_le(x_14, x_14); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +lean_dec(x_14); +lean_dec(x_5); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; +} +else +{ +size_t x_43; size_t x_44; lean_object* x_45; lean_object* x_46; +x_43 = 0; +x_44 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_45 = lean_box(0); +x_46 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3(x_1, x_43, x_44, x_45, x_4, x_5, x_6, x_7, x_8, x_39); +return x_46; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; uint8_t x_8; lean_object* x_9; +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1(x_1, x_2, x_3, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; +x_13 = lean_st_mk_ref(x_12, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_7); +x_16 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed(x_1, x_2, x_3, x_14, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_get(x_14, x_17); +lean_dec(x_14); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_21 = l_Lean_Meta_removeUnused(x_1, x_19, x_7, x_8, x_9, x_10, 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; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_apply_1(x_4, x_27); +x_29 = l_Lean_Meta_withLCtx___at_Lean_Elab_Term_elabSyntheticHole___spec__4___rarg(x_25, x_26, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_21); +if (x_30 == 0) +{ +return x_21; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_21, 0); +x_32 = lean_ctor_get(x_21, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_21); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +uint8_t x_34; +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_16); +if (x_34 == 0) +{ +return x_16; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_16, 0); +x_36 = lean_ctor_get(x_16, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_16); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg___boxed), 11, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_12; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("deprecated", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("oldSectionVars", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("re-enable deprecated behavior of including exactly the section variables used in a declaration", 94, 94); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 0; +x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__1; +x_2 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__1; +x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1; +x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6; +x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_nat_dec_le(x_7, x_6); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_5, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_dec(x_9); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_28 = lean_nat_dec_lt(x_6, x_4); +x_29 = lean_nat_sub(x_1, x_4); +x_30 = lean_nat_add(x_29, x_6); +lean_dec(x_29); +x_31 = lean_array_get_size(x_2); +x_32 = lean_nat_dec_lt(x_30, x_31); +lean_dec(x_31); +if (x_28 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = l_Lean_instInhabitedSyntax; +x_34 = l_outOfBounds___rarg(x_33); +if (x_32 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_30); +x_35 = l_Lean_instInhabitedExpr; +x_36 = l_outOfBounds___rarg(x_35); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_37 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_36, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_39; +x_23 = x_38; +goto block_27; +} +else +{ +uint8_t x_40; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) +{ +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 0); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_array_fget(x_2, x_30); +lean_dec(x_30); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_45 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_47; +x_23 = x_46; +goto block_27; +} +else +{ +uint8_t x_48; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +return x_45; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_45); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +} +else +{ +lean_object* x_52; +x_52 = lean_array_fget(x_3, x_6); +if (x_32 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_30); +x_53 = l_Lean_instInhabitedExpr; +x_54 = l_outOfBounds___rarg(x_53); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_55 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_54, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_57; +x_23 = x_56; +goto block_27; +} +else +{ +uint8_t x_58; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_58 = !lean_is_exclusive(x_55); +if (x_58 == 0) +{ +return x_55; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_55, 0); +x_60 = lean_ctor_get(x_55, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_55); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +else +{ +lean_object* x_62; lean_object* x_63; +x_62 = lean_array_fget(x_2, x_30); +lean_dec(x_30); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_63 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_62, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_65; +x_23 = x_64; +goto block_27; +} +else +{ +uint8_t x_66; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_66 = !lean_is_exclusive(x_63); +if (x_66 == 0) +{ +return x_63; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_63, 0); +x_68 = lean_ctor_get(x_63, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_63); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +block_27: +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_21; +x_6 = x_25; +x_9 = x_24; +x_16 = x_23; +goto _start; +} +} +else +{ +lean_object* x_70; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_9); +lean_ctor_set(x_70, 1, x_16); +return x_70; +} +} +else +{ +lean_object* x_71; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_9); +lean_ctor_set(x_71, 1, x_16); +return x_71; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_array_uget(x_2, x_3); +x_14 = l_Lean_Expr_fvarId_x21(x_13); +lean_dec(x_13); +lean_inc(x_7); +x_15 = l_Lean_FVarId_getType(x_14, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_Lean_Expr_fvarId_x21(x_1); +x_20 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_19, x_17); +lean_dec(x_17); +lean_dec(x_19); +if (x_20 == 0) +{ +size_t x_21; size_t x_22; +lean_free_object(x_15); +x_21 = 1; +x_22 = lean_usize_add(x_3, x_21); +x_3 = x_22; +x_11 = x_18; +goto _start; +} +else +{ +uint8_t x_24; lean_object* x_25; +lean_dec(x_7); +x_24 = 1; +x_25 = lean_box(x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = l_Lean_Expr_fvarId_x21(x_1); +x_29 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_28, x_26); +lean_dec(x_26); +lean_dec(x_28); +if (x_29 == 0) +{ +size_t x_30; size_t x_31; +x_30 = 1; +x_31 = lean_usize_add(x_3, x_30); +x_3 = x_31; +x_11 = x_27; +goto _start; +} +else +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_7); +x_33 = 1; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_27); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_7); +x_36 = !lean_is_exclusive(x_15); +if (x_36 == 0) +{ +return x_15; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_15, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_15); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_7); +x_40 = 0; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_11); +return x_42; +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("included section variable '", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' is not used in '", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("', consider excluding it", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___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, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_8, x_7); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_27; lean_object* x_28; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_9); +x_19 = lean_array_uget(x_6, x_8); +x_91 = lean_array_get_size(x_3); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_nat_dec_lt(x_92, x_91); +if (x_93 == 0) +{ +uint8_t x_94; +lean_dec(x_91); +x_94 = 0; +x_27 = x_94; +x_28 = x_16; +goto block_90; +} +else +{ +size_t x_95; size_t x_96; lean_object* x_97; +x_95 = 0; +x_96 = lean_usize_of_nat(x_91); +lean_dec(x_91); +lean_inc(x_12); +x_97 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(x_19, x_3, x_95, x_96, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = lean_unbox(x_98); +lean_dec(x_98); +x_27 = x_100; +x_28 = x_99; +goto block_90; +} +else +{ +uint8_t x_101; +lean_dec(x_19); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_101 = !lean_is_exclusive(x_97); +if (x_101 == 0) +{ +return x_97; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_97, 0); +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_97); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; +} +} +} +block_26: +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = 1; +x_24 = lean_usize_add(x_8, x_23); +x_8 = x_24; +x_9 = x_22; +x_16 = x_21; +goto _start; +} +block_90: +{ +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Expr_fvarId_x21(x_19); +x_30 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_29, x_4); +if (x_30 == 0) +{ +uint8_t x_31; +x_31 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_29, x_5); +if (x_31 == 0) +{ +if (x_27 == 0) +{ +lean_object* x_32; +lean_inc(x_12); +x_32 = l_Lean_FVarId_getDecl(x_29, x_12, x_13, x_14, x_15, x_28); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_LocalDecl_userName(x_33); +x_36 = l_Lean_Name_hasMacroScopes(x_35); +lean_dec(x_35); +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_38 = l_Lean_MessageData_ofName(x_2); +if (x_36 == 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; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_33); +x_39 = l_Lean_MessageData_ofExpr(x_19); +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_38); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = 1; +lean_inc(x_14); +x_48 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_50; +x_21 = x_49; +goto block_26; +} +else +{ +uint8_t x_51; uint8_t x_52; +x_51 = l_Lean_LocalDecl_binderInfo(x_33); +x_52 = l_Lean_BinderInfo_isInstImplicit(x_51); +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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_33); +x_53 = l_Lean_MessageData_ofExpr(x_19); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_38); +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = 1; +lean_inc(x_14); +x_62 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_60, x_61, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_64; +x_21 = x_63; +goto block_26; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_19); +x_65 = l_Lean_LocalDecl_type(x_33); +lean_dec(x_33); +x_66 = l_Lean_MessageData_ofExpr(x_65); +x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8; +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10; +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = lean_alloc_ctor(6, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_38); +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = 1; +lean_inc(x_14); +x_80 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_78, x_79, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_82; +x_21 = x_81; +goto block_26; +} +} +} +else +{ +uint8_t x_83; +lean_dec(x_19); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_83 = !lean_is_exclusive(x_32); +if (x_83 == 0) +{ +return x_32; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_32, 0); +x_85 = lean_ctor_get(x_32, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_32); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +} +else +{ +lean_object* x_87; +lean_dec(x_29); +lean_dec(x_19); +x_87 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_87; +x_21 = x_28; +goto block_26; +} +} +else +{ +lean_object* x_88; +lean_dec(x_29); +lean_dec(x_19); +x_88 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_88; +x_21 = x_28; +goto block_26; +} +} +else +{ +lean_object* x_89; +lean_dec(x_29); +lean_dec(x_19); +x_89 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_89; +x_21 = x_28; +goto block_26; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_nat_dec_le(x_7, x_6); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_5, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_dec(x_9); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_28 = lean_nat_dec_lt(x_6, x_4); +x_29 = lean_nat_sub(x_1, x_4); +x_30 = lean_nat_add(x_29, x_6); +lean_dec(x_29); +x_31 = lean_array_get_size(x_2); +x_32 = lean_nat_dec_lt(x_30, x_31); +lean_dec(x_31); +if (x_28 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = l_Lean_instInhabitedSyntax; +x_34 = l_outOfBounds___rarg(x_33); +if (x_32 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_30); +x_35 = l_Lean_instInhabitedExpr; +x_36 = l_outOfBounds___rarg(x_35); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_37 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_36, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_39; +x_23 = x_38; +goto block_27; +} +else +{ +uint8_t x_40; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) +{ +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 0); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_array_fget(x_2, x_30); +lean_dec(x_30); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_45 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_47; +x_23 = x_46; +goto block_27; +} +else +{ +uint8_t x_48; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +return x_45; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_45); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +} +else +{ +lean_object* x_52; +x_52 = lean_array_fget(x_3, x_6); +if (x_32 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_30); +x_53 = l_Lean_instInhabitedExpr; +x_54 = l_outOfBounds___rarg(x_53); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_55 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_54, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_57; +x_23 = x_56; +goto block_27; +} +else +{ +uint8_t x_58; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_58 = !lean_is_exclusive(x_55); +if (x_58 == 0) +{ +return x_55; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_55, 0); +x_60 = lean_ctor_get(x_55, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_55); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +else +{ +lean_object* x_62; lean_object* x_63; +x_62 = lean_array_fget(x_2, x_30); +lean_dec(x_30); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_63 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_62, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_22 = x_65; +x_23 = x_64; +goto block_27; +} +else +{ +uint8_t x_66; +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +x_66 = !lean_is_exclusive(x_63); +if (x_66 == 0) +{ +return x_63; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_63, 0); +x_68 = lean_ctor_get(x_63, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_63); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +block_27: +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_21; +x_6 = x_25; +x_9 = x_24; +x_16 = x_23; +goto _start; +} +} +else +{ +lean_object* x_70; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_9); +lean_ctor_set(x_70, 1, x_16); +return x_70; +} +} +else +{ +lean_object* x_71; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_9); +lean_ctor_set(x_71, 1, x_16); +return x_71; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_eq(x_3, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_array_uget(x_2, x_3); +x_14 = l_Lean_Expr_fvarId_x21(x_13); +lean_dec(x_13); +lean_inc(x_7); +x_15 = l_Lean_FVarId_getType(x_14, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_Lean_Expr_fvarId_x21(x_1); +x_20 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_19, x_17); +lean_dec(x_17); +lean_dec(x_19); +if (x_20 == 0) +{ +size_t x_21; size_t x_22; +lean_free_object(x_15); +x_21 = 1; +x_22 = lean_usize_add(x_3, x_21); +x_3 = x_22; +x_11 = x_18; +goto _start; +} +else +{ +uint8_t x_24; lean_object* x_25; +lean_dec(x_7); +x_24 = 1; +x_25 = lean_box(x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = l_Lean_Expr_fvarId_x21(x_1); +x_29 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_28, x_26); +lean_dec(x_26); +lean_dec(x_28); +if (x_29 == 0) +{ +size_t x_30; size_t x_31; +x_30 = 1; +x_31 = lean_usize_add(x_3, x_30); +x_3 = x_31; +x_11 = x_27; +goto _start; +} +else +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_7); +x_33 = 1; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_27); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_7); +x_36 = !lean_is_exclusive(x_15); +if (x_36 == 0) +{ +return x_15; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_15, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_15); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_7); +x_40 = 0; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_11); +return x_42; +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_8, x_7); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_27; lean_object* x_28; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_9); +x_19 = lean_array_uget(x_6, x_8); +x_91 = lean_array_get_size(x_3); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_nat_dec_lt(x_92, x_91); +if (x_93 == 0) +{ +uint8_t x_94; +lean_dec(x_91); +x_94 = 0; +x_27 = x_94; +x_28 = x_16; +goto block_90; +} +else +{ +size_t x_95; size_t x_96; lean_object* x_97; +x_95 = 0; +x_96 = lean_usize_of_nat(x_91); +lean_dec(x_91); +lean_inc(x_12); +x_97 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_19, x_3, x_95, x_96, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = lean_unbox(x_98); +lean_dec(x_98); +x_27 = x_100; +x_28 = x_99; +goto block_90; +} +else +{ +uint8_t x_101; +lean_dec(x_19); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_101 = !lean_is_exclusive(x_97); +if (x_101 == 0) +{ +return x_97; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_97, 0); +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_97); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; +} +} +} +block_26: +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = 1; +x_24 = lean_usize_add(x_8, x_23); +x_8 = x_24; +x_9 = x_22; +x_16 = x_21; +goto _start; +} +block_90: +{ +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Expr_fvarId_x21(x_19); +x_30 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_29, x_4); +if (x_30 == 0) +{ +uint8_t x_31; +x_31 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_29, x_5); +if (x_31 == 0) +{ +if (x_27 == 0) +{ +lean_object* x_32; +lean_inc(x_12); +x_32 = l_Lean_FVarId_getDecl(x_29, x_12, x_13, x_14, x_15, x_28); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_LocalDecl_userName(x_33); +x_36 = l_Lean_Name_hasMacroScopes(x_35); +lean_dec(x_35); +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_38 = l_Lean_MessageData_ofName(x_2); +if (x_36 == 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; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_33); +x_39 = l_Lean_MessageData_ofExpr(x_19); +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_38); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = 1; +lean_inc(x_14); +x_48 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_50; +x_21 = x_49; +goto block_26; +} +else +{ +uint8_t x_51; uint8_t x_52; +x_51 = l_Lean_LocalDecl_binderInfo(x_33); +x_52 = l_Lean_BinderInfo_isInstImplicit(x_51); +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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_33); +x_53 = l_Lean_MessageData_ofExpr(x_19); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_38); +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = 1; +lean_inc(x_14); +x_62 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_60, x_61, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_64; +x_21 = x_63; +goto block_26; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_19); +x_65 = l_Lean_LocalDecl_type(x_33); +lean_dec(x_33); +x_66 = l_Lean_MessageData_ofExpr(x_65); +x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8; +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10; +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = lean_alloc_ctor(6, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_38); +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = 1; +lean_inc(x_14); +x_80 = l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(x_37, x_78, x_79, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_82; +x_21 = x_81; +goto block_26; +} +} +} +else +{ +uint8_t x_83; +lean_dec(x_19); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +x_83 = !lean_is_exclusive(x_32); +if (x_83 == 0) +{ +return x_32; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_32, 0); +x_85 = lean_ctor_get(x_32, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_32); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +} +else +{ +lean_object* x_87; +lean_dec(x_29); +lean_dec(x_19); +x_87 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_87; +x_21 = x_28; +goto block_26; +} +} +else +{ +lean_object* x_88; +lean_dec(x_29); +lean_dec(x_19); +x_88 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_88; +x_21 = x_28; +goto block_26; +} +} +else +{ +lean_object* x_89; +lean_dec(x_29); +lean_dec(x_19); +x_89 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; +x_20 = x_89; +x_21 = x_28; +goto block_26; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_ctor_get_uint8(x_7, sizeof(void*)*12 + 1); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_7); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_7, 5); +lean_dec(x_12); +x_13 = 0; +lean_ctor_set(x_7, 5, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*12 + 1, x_13); +x_14 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_15 = lean_ctor_get(x_7, 0); +x_16 = lean_ctor_get(x_7, 1); +x_17 = lean_ctor_get(x_7, 2); +x_18 = lean_ctor_get(x_7, 3); +x_19 = lean_ctor_get(x_7, 4); +x_20 = lean_ctor_get(x_7, 6); +x_21 = lean_ctor_get(x_7, 7); +x_22 = lean_ctor_get(x_7, 8); +x_23 = lean_ctor_get(x_7, 9); +x_24 = lean_ctor_get(x_7, 10); +x_25 = lean_ctor_get_uint8(x_7, sizeof(void*)*12); +x_26 = lean_ctor_get(x_7, 11); +lean_inc(x_26); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_7); +x_27 = 0; +x_28 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_28, 0, x_15); +lean_ctor_set(x_28, 1, x_16); +lean_ctor_set(x_28, 2, x_17); +lean_ctor_set(x_28, 3, x_18); +lean_ctor_set(x_28, 4, x_19); +lean_ctor_set(x_28, 5, x_1); +lean_ctor_set(x_28, 6, x_20); +lean_ctor_set(x_28, 7, x_21); +lean_ctor_set(x_28, 8, x_22); +lean_ctor_set(x_28, 9, x_23); +lean_ctor_set(x_28, 10, x_24); +lean_ctor_set(x_28, 11, x_26); +lean_ctor_set_uint8(x_28, sizeof(void*)*12, x_25); +lean_ctor_set_uint8(x_28, sizeof(void*)*12 + 1, x_27); +x_29 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_28, x_8, x_9); +return x_29; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_7, x_6); -if (x_17 == 0) +else { -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_eq(x_5, x_18); -if (x_19 == 0) +uint8_t x_30; +x_30 = !lean_is_exclusive(x_7); +if (x_30 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -lean_dec(x_9); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_5, x_20); -lean_dec(x_5); -x_28 = lean_nat_dec_lt(x_6, x_4); -x_29 = lean_nat_sub(x_1, x_4); -x_30 = lean_nat_add(x_29, x_6); -lean_dec(x_29); -x_31 = lean_array_get_size(x_2); -x_32 = lean_nat_dec_lt(x_30, x_31); +lean_object* x_31; uint8_t x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_7, 5); lean_dec(x_31); -if (x_28 == 0) +lean_inc(x_1); +x_32 = l_Lean_Syntax_hasMissing(x_1); +lean_ctor_set(x_7, 5, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*12 + 1, x_32); +x_33 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_33; +} +else { -lean_object* x_33; lean_object* x_34; -x_33 = l_Lean_instInhabitedSyntax; -x_34 = l_outOfBounds___rarg(x_33); -if (x_32 == 0) +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_34 = lean_ctor_get(x_7, 0); +x_35 = lean_ctor_get(x_7, 1); +x_36 = lean_ctor_get(x_7, 2); +x_37 = lean_ctor_get(x_7, 3); +x_38 = lean_ctor_get(x_7, 4); +x_39 = lean_ctor_get(x_7, 6); +x_40 = lean_ctor_get(x_7, 7); +x_41 = lean_ctor_get(x_7, 8); +x_42 = lean_ctor_get(x_7, 9); +x_43 = lean_ctor_get(x_7, 10); +x_44 = lean_ctor_get_uint8(x_7, sizeof(void*)*12); +x_45 = lean_ctor_get(x_7, 11); +lean_inc(x_45); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_7); +lean_inc(x_1); +x_46 = l_Lean_Syntax_hasMissing(x_1); +x_47 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_47, 0, x_34); +lean_ctor_set(x_47, 1, x_35); +lean_ctor_set(x_47, 2, x_36); +lean_ctor_set(x_47, 3, x_37); +lean_ctor_set(x_47, 4, x_38); +lean_ctor_set(x_47, 5, x_1); +lean_ctor_set(x_47, 6, x_39); +lean_ctor_set(x_47, 7, x_40); +lean_ctor_set(x_47, 8, x_41); +lean_ctor_set(x_47, 9, x_42); +lean_ctor_set(x_47, 10, x_43); +lean_ctor_set(x_47, 11, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*12, x_44); +lean_ctor_set_uint8(x_47, sizeof(void*)*12 + 1, x_46); +x_48 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_47, x_8, x_9); +return x_48; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_30); -x_35 = l_Lean_instInhabitedExpr; -x_36 = l_outOfBounds___rarg(x_35); +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_1, 3); +x_19 = lean_array_get_size(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_box(0); +lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -x_37 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_36, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_37) == 0) +lean_inc(x_19); +x_23 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(x_2, x_9, x_18, x_19, x_19, x_20, x_19, x_21, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_19); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; -x_22 = x_39; -x_23 = x_38; -goto block_27; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_10); +x_26 = lean_box(0); +x_27 = lean_ctor_get(x_11, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +x_29 = lean_ctor_get(x_11, 2); +lean_inc(x_29); +x_30 = lean_ctor_get_uint8(x_11, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_11, 3); +lean_inc(x_33); +x_34 = lean_ctor_get(x_11, 4); +lean_inc(x_34); +x_35 = lean_ctor_get(x_11, 5); +lean_inc(x_35); +x_36 = lean_ctor_get(x_11, 6); +lean_inc(x_36); +x_37 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_11, 7); +lean_inc(x_42); +x_43 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 9); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_3); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = 1; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_45); +x_47 = l_Lean_Elab_Term_elabTermEnsuringType(x_4, x_25, x_46, x_46, x_26, x_45, x_12, x_13, x_14, x_15, x_16, x_24); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = 1; +x_51 = 0; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_52 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_50, x_51, x_45, x_12, x_13, x_14, x_15, x_16, x_49); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_54 = l_Lean_Elab_Term_instantiateMVarsProfiling(x_48, x_13, x_14, x_15, x_16, x_53); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = 1; +x_58 = l_Lean_Meta_mkLambdaFVars(x_9, x_55, x_51, x_46, x_51, x_57, x_13, x_14, x_15, x_16, x_56); +if (lean_obj_tag(x_58) == 0) +{ +uint8_t x_59; +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = lean_ctor_get(x_58, 0); +x_61 = lean_ctor_get(x_58, 1); +x_62 = l_Lean_Expr_hasSorry(x_5); +if (x_62 == 0) +{ +uint8_t x_63; +x_63 = l_Lean_Expr_hasSorry(x_60); +if (x_63 == 0) +{ +size_t x_64; size_t x_65; lean_object* x_66; +lean_free_object(x_58); +x_64 = lean_array_size(x_6); +x_65 = 0; +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_7, x_8, x_6, x_5, x_60, x_6, x_64, x_65, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_61); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +x_68 = lean_ctor_get(x_66, 0); +lean_dec(x_68); +lean_ctor_set(x_66, 0, x_60); +return x_66; } else { -uint8_t x_40; -lean_dec(x_21); +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_60); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +else +{ +uint8_t x_71; +lean_dec(x_60); +x_71 = !lean_is_exclusive(x_66); +if (x_71 == 0) +{ +return x_66; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_66, 0); +x_73 = lean_ctor_get(x_66, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_66); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +else +{ +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -x_40 = !lean_is_exclusive(x_37); -if (x_40 == 0) +lean_dec(x_8); +return x_58; +} +} +else { -return x_37; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +return x_58; +} } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_ctor_get(x_37, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_37); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_75 = lean_ctor_get(x_58, 0); +x_76 = lean_ctor_get(x_58, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_58); +x_77 = l_Lean_Expr_hasSorry(x_5); +if (x_77 == 0) +{ +uint8_t x_78; +x_78 = l_Lean_Expr_hasSorry(x_75); +if (x_78 == 0) +{ +size_t x_79; size_t x_80; lean_object* x_81; +x_79 = lean_array_size(x_6); +x_80 = 0; +x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_7, x_8, x_6, x_5, x_75, x_6, x_79, x_80, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); } +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(0, 2, 0); +} else { + x_84 = x_83; } +lean_ctor_set(x_84, 0, x_75); +lean_ctor_set(x_84, 1, x_82); +return x_84; } else { -lean_object* x_44; lean_object* x_45; -x_44 = lean_array_fget(x_2, x_30); -lean_dec(x_30); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_45 = l_Lean_Elab_Term_addLocalVarInfo(x_34, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_45) == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_75); +x_85 = lean_ctor_get(x_81, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_81, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_87 = x_81; +} else { + lean_dec_ref(x_81); + x_87 = lean_box(0); +} +if (lean_is_scalar(x_87)) { + x_88 = lean_alloc_ctor(1, 2, 0); +} else { + x_88 = x_87; +} +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_86); +return x_88; +} +} +else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -lean_dec(x_45); -x_47 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; -x_22 = x_47; -x_23 = x_46; -goto block_27; +lean_object* x_89; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_75); +lean_ctor_set(x_89, 1, x_76); +return x_89; +} } else { -uint8_t x_48; -lean_dec(x_21); +lean_object* x_90; +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -x_48 = !lean_is_exclusive(x_45); -if (x_48 == 0) +lean_dec(x_8); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_75); +lean_ctor_set(x_90, 1, x_76); +return x_90; +} +} +} +else { -return x_45; +uint8_t x_91; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +x_91 = !lean_is_exclusive(x_58); +if (x_91 == 0) +{ +return x_58; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_45, 0); -x_50 = lean_ctor_get(x_45, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_45); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_58, 0); +x_93 = lean_ctor_get(x_58, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_58); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; } } } +else +{ +uint8_t x_95; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +x_95 = !lean_is_exclusive(x_54); +if (x_95 == 0) +{ +return x_54; } else { -lean_object* x_52; -x_52 = lean_array_fget(x_3, x_6); -if (x_32 == 0) +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_54, 0); +x_97 = lean_ctor_get(x_54, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_54); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_30); -x_53 = l_Lean_instInhabitedExpr; -x_54 = l_outOfBounds___rarg(x_53); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_55 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_54, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_55) == 0) +uint8_t x_99; +lean_dec(x_48); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +x_99 = !lean_is_exclusive(x_52); +if (x_99 == 0) { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; -x_22 = x_57; -x_23 = x_56; -goto block_27; +return x_52; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_52, 0); +x_101 = lean_ctor_get(x_52, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_52); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; +} +} +} +else +{ +uint8_t x_103; +lean_dec(x_45); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +x_103 = !lean_is_exclusive(x_47); +if (x_103 == 0) +{ +return x_47; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_47, 0); +x_105 = lean_ctor_get(x_47, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_47); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} } else { -uint8_t x_58; -lean_dec(x_21); +uint8_t x_107; +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_6); -x_58 = !lean_is_exclusive(x_55); -if (x_58 == 0) +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +x_107 = !lean_is_exclusive(x_23); +if (x_107 == 0) { -return x_55; +return x_23; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_55, 0); -x_60 = lean_ctor_get(x_55, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_55); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_23, 0); +x_109 = lean_ctor_get(x_23, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_23); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: { -lean_object* x_62; lean_object* x_63; -x_62 = lean_array_fget(x_2, x_30); -lean_dec(x_30); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_1, 3); +x_19 = lean_array_get_size(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_box(0); +lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -x_63 = l_Lean_Elab_Term_addLocalVarInfo(x_52, x_62, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_63) == 0) +lean_inc(x_19); +x_23 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(x_2, x_9, x_18, x_19, x_19, x_20, x_19, x_21, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_19); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -x_65 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; -x_22 = x_65; -x_23 = x_64; -goto block_27; -} -else +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_10); +x_26 = lean_box(0); +x_27 = lean_ctor_get(x_11, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +x_29 = lean_ctor_get(x_11, 2); +lean_inc(x_29); +x_30 = lean_ctor_get_uint8(x_11, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_11, 3); +lean_inc(x_33); +x_34 = lean_ctor_get(x_11, 4); +lean_inc(x_34); +x_35 = lean_ctor_get(x_11, 5); +lean_inc(x_35); +x_36 = lean_ctor_get(x_11, 6); +lean_inc(x_36); +x_37 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_11, 7); +lean_inc(x_42); +x_43 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_11, sizeof(void*)*9 + 9); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_3); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = 1; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_45); +x_47 = l_Lean_Elab_Term_elabTermEnsuringType(x_4, x_25, x_46, x_46, x_26, x_45, x_12, x_13, x_14, x_15, x_16, x_24); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_66; -lean_dec(x_21); -lean_dec(x_15); +lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = 1; +x_51 = 0; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_52 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_50, x_51, x_45, x_12, x_13, x_14, x_15, x_16, x_49); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_54 = l_Lean_Elab_Term_instantiateMVarsProfiling(x_48, x_13, x_14, x_15, x_16, x_53); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = 1; +x_58 = l_Lean_Meta_mkLambdaFVars(x_9, x_55, x_51, x_46, x_51, x_57, x_13, x_14, x_15, x_16, x_56); +if (lean_obj_tag(x_58) == 0) +{ +uint8_t x_59; +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = lean_ctor_get(x_58, 0); +x_61 = lean_ctor_get(x_58, 1); +x_62 = l_Lean_Expr_hasSorry(x_5); +if (x_62 == 0) +{ +uint8_t x_63; +x_63 = l_Lean_Expr_hasSorry(x_60); +if (x_63 == 0) +{ +size_t x_64; size_t x_65; lean_object* x_66; +lean_free_object(x_58); +x_64 = lean_array_size(x_6); +x_65 = 0; +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_7, x_8, x_6, x_5, x_60, x_6, x_64, x_65, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_61); +lean_dec(x_16); lean_dec(x_14); -lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -x_66 = !lean_is_exclusive(x_63); -if (x_66 == 0) +if (lean_obj_tag(x_66) == 0) { -return x_63; +uint8_t x_67; +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +x_68 = lean_ctor_get(x_66, 0); +lean_dec(x_68); +lean_ctor_set(x_66, 0, x_60); +return x_66; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_63, 0); -x_68 = lean_ctor_get(x_63, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_63); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -return x_69; -} +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_60); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } +else +{ +uint8_t x_71; +lean_dec(x_60); +x_71 = !lean_is_exclusive(x_66); +if (x_71 == 0) +{ +return x_66; } -block_27: +else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_5 = x_21; -x_6 = x_25; -x_9 = x_24; -x_16 = x_23; -goto _start; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_66, 0); +x_73 = lean_ctor_get(x_66, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_66); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} } } else { -lean_object* x_70; +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_9); -lean_ctor_set(x_70, 1, x_16); -return x_70; +lean_dec(x_8); +return x_58; } } else { -lean_object* x_71; +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_9); -lean_ctor_set(x_71, 1, x_16); -return x_71; -} +lean_dec(x_8); +return x_58; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -uint8_t x_10; -x_10 = lean_ctor_get_uint8(x_7, sizeof(void*)*12 + 1); -if (x_10 == 0) +lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_75 = lean_ctor_get(x_58, 0); +x_76 = lean_ctor_get(x_58, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_58); +x_77 = l_Lean_Expr_hasSorry(x_5); +if (x_77 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_7); -if (x_11 == 0) +uint8_t x_78; +x_78 = l_Lean_Expr_hasSorry(x_75); +if (x_78 == 0) { -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_7, 5); +size_t x_79; size_t x_80; lean_object* x_81; +x_79 = lean_array_size(x_6); +x_80 = 0; +x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_7, x_8, x_6, x_5, x_75, x_6, x_79, x_80, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_12); -x_13 = 0; -lean_ctor_set(x_7, 5, x_1); -lean_ctor_set_uint8(x_7, sizeof(void*)*12 + 1, x_13); -x_14 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_14; +lean_dec(x_11); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); +} +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(0, 2, 0); +} else { + x_84 = x_83; +} +lean_ctor_set(x_84, 0, x_75); +lean_ctor_set(x_84, 1, x_82); +return x_84; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_15 = lean_ctor_get(x_7, 0); -x_16 = lean_ctor_get(x_7, 1); -x_17 = lean_ctor_get(x_7, 2); -x_18 = lean_ctor_get(x_7, 3); -x_19 = lean_ctor_get(x_7, 4); -x_20 = lean_ctor_get(x_7, 6); -x_21 = lean_ctor_get(x_7, 7); -x_22 = lean_ctor_get(x_7, 8); -x_23 = lean_ctor_get(x_7, 9); -x_24 = lean_ctor_get(x_7, 10); -x_25 = lean_ctor_get_uint8(x_7, sizeof(void*)*12); -x_26 = lean_ctor_get(x_7, 11); -lean_inc(x_26); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_7); -x_27 = 0; -x_28 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_28, 0, x_15); -lean_ctor_set(x_28, 1, x_16); -lean_ctor_set(x_28, 2, x_17); -lean_ctor_set(x_28, 3, x_18); -lean_ctor_set(x_28, 4, x_19); -lean_ctor_set(x_28, 5, x_1); -lean_ctor_set(x_28, 6, x_20); -lean_ctor_set(x_28, 7, x_21); -lean_ctor_set(x_28, 8, x_22); -lean_ctor_set(x_28, 9, x_23); -lean_ctor_set(x_28, 10, x_24); -lean_ctor_set(x_28, 11, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*12, x_25); -lean_ctor_set_uint8(x_28, sizeof(void*)*12 + 1, x_27); -x_29 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_28, x_8, x_9); -return x_29; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_75); +x_85 = lean_ctor_get(x_81, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_81, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_87 = x_81; +} else { + lean_dec_ref(x_81); + x_87 = lean_box(0); +} +if (lean_is_scalar(x_87)) { + x_88 = lean_alloc_ctor(1, 2, 0); +} else { + x_88 = x_87; +} +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_86); +return x_88; } } else { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_7); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_7, 5); -lean_dec(x_31); -lean_inc(x_1); -x_32 = l_Lean_Syntax_hasMissing(x_1); -lean_ctor_set(x_7, 5, x_1); -lean_ctor_set_uint8(x_7, sizeof(void*)*12 + 1, x_32); -x_33 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_33; +lean_object* x_89; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_75); +lean_ctor_set(x_89, 1, x_76); +return x_89; +} } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -x_34 = lean_ctor_get(x_7, 0); -x_35 = lean_ctor_get(x_7, 1); -x_36 = lean_ctor_get(x_7, 2); -x_37 = lean_ctor_get(x_7, 3); -x_38 = lean_ctor_get(x_7, 4); -x_39 = lean_ctor_get(x_7, 6); -x_40 = lean_ctor_get(x_7, 7); -x_41 = lean_ctor_get(x_7, 8); -x_42 = lean_ctor_get(x_7, 9); -x_43 = lean_ctor_get(x_7, 10); -x_44 = lean_ctor_get_uint8(x_7, sizeof(void*)*12); -x_45 = lean_ctor_get(x_7, 11); -lean_inc(x_45); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_7); -lean_inc(x_1); -x_46 = l_Lean_Syntax_hasMissing(x_1); -x_47 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_47, 0, x_34); -lean_ctor_set(x_47, 1, x_35); -lean_ctor_set(x_47, 2, x_36); -lean_ctor_set(x_47, 3, x_37); -lean_ctor_set(x_47, 4, x_38); -lean_ctor_set(x_47, 5, x_1); -lean_ctor_set(x_47, 6, x_39); -lean_ctor_set(x_47, 7, x_40); -lean_ctor_set(x_47, 8, x_41); -lean_ctor_set(x_47, 9, x_42); -lean_ctor_set(x_47, 10, x_43); -lean_ctor_set(x_47, 11, x_45); -lean_ctor_set_uint8(x_47, sizeof(void*)*12, x_44); -lean_ctor_set_uint8(x_47, sizeof(void*)*12 + 1, x_46); -x_48 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_47, x_8, x_9); -return x_48; -} +lean_object* x_90; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_75); +lean_ctor_set(x_90, 1, x_76); +return x_90; } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_1, 3); -x_15 = lean_array_get_size(x_14); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_box(0); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_15); -x_19 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(x_2, x_5, x_14, x_15, x_15, x_16, x_15, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +uint8_t x_91; +lean_dec(x_16); lean_dec(x_15); -if (lean_obj_tag(x_19) == 0) +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +x_91 = !lean_is_exclusive(x_58); +if (x_91 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; uint8_t x_39; uint8_t x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_6); -x_22 = lean_box(0); -x_23 = lean_ctor_get(x_7, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_7, 1); -lean_inc(x_24); -x_25 = lean_ctor_get(x_7, 2); -lean_inc(x_25); -x_26 = lean_ctor_get_uint8(x_7, sizeof(void*)*9); -x_27 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 1); -x_28 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 2); -x_29 = lean_ctor_get(x_7, 3); -lean_inc(x_29); -x_30 = lean_ctor_get(x_7, 4); -lean_inc(x_30); -x_31 = lean_ctor_get(x_7, 5); -lean_inc(x_31); -x_32 = lean_ctor_get(x_7, 6); -lean_inc(x_32); -x_33 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 3); -x_34 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 4); -x_35 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 5); -x_36 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 6); -x_37 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 7); -x_38 = lean_ctor_get(x_7, 7); -lean_inc(x_38); -x_39 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 8); -x_40 = lean_ctor_get_uint8(x_7, sizeof(void*)*9 + 9); -x_41 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_24); -lean_ctor_set(x_41, 2, x_25); -lean_ctor_set(x_41, 3, x_29); -lean_ctor_set(x_41, 4, x_30); -lean_ctor_set(x_41, 5, x_31); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_38); -lean_ctor_set(x_41, 8, x_3); -lean_ctor_set_uint8(x_41, sizeof(void*)*9, x_26); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 1, x_27); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 2, x_28); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 3, x_33); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 4, x_34); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 5, x_35); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 6, x_36); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 7, x_37); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 8, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*9 + 9, x_40); -x_42 = 1; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_41); -x_43 = l_Lean_Elab_Term_elabTermEnsuringType(x_4, x_21, x_42, x_42, x_22, x_41, x_8, x_9, x_10, x_11, x_12, x_20); -if (lean_obj_tag(x_43) == 0) +return x_58; +} +else { -lean_object* x_44; lean_object* x_45; uint8_t x_46; uint8_t x_47; lean_object* x_48; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = 1; -x_47 = 0; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_48 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_46, x_47, x_41, x_8, x_9, x_10, x_11, x_12, x_45); -if (lean_obj_tag(x_48) == 0) +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_58, 0); +x_93 = lean_ctor_get(x_58, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_58); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_44, x_7, x_8, x_9, x_10, x_11, x_12, x_49); -lean_dec(x_8); -lean_dec(x_7); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = l_Lean_Meta_mkLambdaFVars(x_5, x_51, x_47, x_42, x_47, x_53, x_9, x_10, x_11, x_12, x_52); +uint8_t x_95; +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); +x_95 = !lean_is_exclusive(x_54); +if (x_95 == 0) +{ return x_54; } else { -uint8_t x_55; -lean_dec(x_44); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_54, 0); +x_97 = lean_ctor_get(x_54, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_54); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_48); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_55 = !lean_is_exclusive(x_48); -if (x_55 == 0) +x_99 = !lean_is_exclusive(x_52); +if (x_99 == 0) { -return x_48; +return x_52; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_48, 0); -x_57 = lean_ctor_get(x_48, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_48); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_52, 0); +x_101 = lean_ctor_get(x_52, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_52); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; } } } else { -uint8_t x_59; -lean_dec(x_41); +uint8_t x_103; +lean_dec(x_45); +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_5); -x_59 = !lean_is_exclusive(x_43); -if (x_59 == 0) +x_103 = !lean_is_exclusive(x_47); +if (x_103 == 0) { -return x_43; +return x_47; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_43, 0); -x_61 = lean_ctor_get(x_43, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_43); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_47, 0); +x_105 = lean_ctor_get(x_47, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_47); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } } else { -uint8_t x_63; +uint8_t x_107; +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_63 = !lean_is_exclusive(x_19); -if (x_63 == 0) +x_107 = !lean_is_exclusive(x_23); +if (x_107 == 0) { -return x_19; +return x_23; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_19, 0); -x_65 = lean_ctor_get(x_19, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_19); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_23, 0); +x_109 = lean_ctor_get(x_23, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_23); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; -lean_inc(x_8); -lean_inc(x_4); -x_11 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_2, 5); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_1, 5); lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 4); +x_15 = lean_ctor_get(x_1, 4); lean_inc(x_15); lean_inc(x_15); x_16 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1___boxed), 13, 4); -lean_closure_set(x_17, 0, x_2); +lean_inc(x_14); +x_17 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3___boxed), 17, 8); +lean_closure_set(x_17, 0, x_1); lean_closure_set(x_17, 1, x_15); -lean_closure_set(x_17, 2, x_3); -lean_closure_set(x_17, 3, x_12); +lean_closure_set(x_17, 2, x_2); +lean_closure_set(x_17, 3, x_3); +lean_closure_set(x_17, 4, x_14); +lean_closure_set(x_17, 5, x_6); +lean_closure_set(x_17, 6, x_4); +lean_closure_set(x_17, 7, x_5); x_18 = 0; -x_19 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg(x_14, x_16, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_19 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg(x_14, x_16, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_19; } +} +static lean_object* _init_l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Term_deprecated_oldSectionVars; +return x_1; +} +} +static lean_object* _init_l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +lean_inc(x_13); +lean_inc(x_9); +x_16 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_28; uint8_t x_29; uint8_t x_30; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_28 = lean_ctor_get(x_13, 2); +lean_inc(x_28); +x_29 = lean_ctor_get_uint8(x_4, sizeof(void*)*9); +x_30 = l_Lean_Elab_DefKind_isTheorem(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +lean_dec(x_28); +lean_dec(x_6); +x_31 = lean_box(0); +x_19 = x_31; +goto block_27; +} else { -uint8_t x_20; +lean_object* x_32; uint8_t x_33; +x_32 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1; +x_33 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_32); +lean_dec(x_28); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__4), 13, 5); +lean_closure_set(x_34, 0, x_2); +lean_closure_set(x_34, 1, x_3); +lean_closure_set(x_34, 2, x_17); +lean_closure_set(x_34, 3, x_4); +lean_closure_set(x_34, 4, x_5); +x_35 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2; +x_36 = lean_array_push(x_35, x_6); +x_37 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(x_7, x_8, x_36, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +lean_dec(x_36); +return x_37; +} +else +{ +lean_object* x_38; +lean_dec(x_6); +x_38 = lean_box(0); +x_19 = x_38; +goto block_27; +} +} +block_27: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +lean_dec(x_19); +x_20 = lean_ctor_get(x_2, 5); +lean_inc(x_20); +x_21 = lean_ctor_get(x_2, 4); +lean_inc(x_21); +lean_inc(x_21); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +lean_inc(x_20); +x_24 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___boxed), 17, 8); +lean_closure_set(x_24, 0, x_2); +lean_closure_set(x_24, 1, x_21); +lean_closure_set(x_24, 2, x_3); +lean_closure_set(x_24, 3, x_17); +lean_closure_set(x_24, 4, x_20); +lean_closure_set(x_24, 5, x_23); +lean_closure_set(x_24, 6, x_4); +lean_closure_set(x_24, 7, x_5); +x_25 = 0; +x_26 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg(x_20, x_22, x_24, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +return x_26; +} +} +else +{ +uint8_t x_39; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_11); -if (x_20 == 0) +x_39 = !lean_is_exclusive(x_16); +if (x_39 == 0) { -return x_11; +return x_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_11, 0); -x_22 = lean_ctor_get(x_11, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_11); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_16, 0); +x_41 = lean_ctor_get(x_16, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_16); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; uint8_t x_12; @@ -18720,7 +21853,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_11, 1); lean_inc(x_9); -x_15 = l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_14); +x_15 = l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_14); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -18845,7 +21978,7 @@ lean_inc(x_43); lean_inc(x_42); lean_dec(x_11); lean_inc(x_9); -x_44 = l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_43); +x_44 = l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_43); if (lean_obj_tag(x_44) == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_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; @@ -18932,102 +22065,106 @@ return x_62; } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 6); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_3, 2); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2), 10, 3); -lean_closure_set(x_18, 0, x_4); -lean_closure_set(x_18, 1, x_3); -lean_closure_set(x_18, 2, x_2); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); -lean_closure_set(x_19, 0, x_17); -lean_closure_set(x_19, 1, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); -lean_closure_set(x_20, 0, x_16); -lean_closure_set(x_20, 1, x_19); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_ctor_get(x_5, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 2); +lean_inc(x_19); +lean_inc(x_7); +x_20 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed), 15, 8); +lean_closure_set(x_20, 0, x_8); +lean_closure_set(x_20, 1, x_6); +lean_closure_set(x_20, 2, x_4); +lean_closure_set(x_20, 3, x_5); +lean_closure_set(x_20, 4, x_7); +lean_closure_set(x_20, 5, x_3); +lean_closure_set(x_20, 6, x_1); +lean_closure_set(x_20, 7, x_2); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); +lean_closure_set(x_21, 0, x_19); +lean_closure_set(x_21, 1, x_20); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); +lean_closure_set(x_22, 0, x_7); +lean_closure_set(x_22, 1, x_21); +x_23 = !lean_is_exclusive(x_10); +if (x_23 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 8); -lean_dec(x_22); -lean_ctor_set(x_6, 8, x_5); -x_23 = lean_box(0); -x_24 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_6, x_7, x_8, x_9, x_23, x_11, x_12, x_13); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_10, 8); +lean_dec(x_24); +lean_ctor_set(x_10, 8, x_9); +x_25 = lean_box(0); +x_26 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_10, x_11, x_12, x_13, x_25, x_15, x_16, x_17); +return x_26; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_6, 0); -x_26 = lean_ctor_get(x_6, 1); -x_27 = lean_ctor_get(x_6, 2); -x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); -x_29 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); -x_30 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); -x_31 = lean_ctor_get(x_6, 3); -x_32 = lean_ctor_get(x_6, 4); -x_33 = lean_ctor_get(x_6, 5); -x_34 = lean_ctor_get(x_6, 6); -x_35 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); -x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); -x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); -x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); -x_39 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); -x_40 = lean_ctor_get(x_6, 7); -x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); -x_42 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); -lean_inc(x_40); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_27 = lean_ctor_get(x_10, 0); +x_28 = lean_ctor_get(x_10, 1); +x_29 = lean_ctor_get(x_10, 2); +x_30 = lean_ctor_get_uint8(x_10, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_10, 3); +x_34 = lean_ctor_get(x_10, 4); +x_35 = lean_ctor_get(x_10, 5); +x_36 = lean_ctor_get(x_10, 6); +x_37 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_10, 7); +x_43 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 9); +lean_inc(x_42); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); +lean_inc(x_29); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_6); -x_43 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_43, 0, x_25); -lean_ctor_set(x_43, 1, x_26); -lean_ctor_set(x_43, 2, x_27); -lean_ctor_set(x_43, 3, x_31); -lean_ctor_set(x_43, 4, x_32); -lean_ctor_set(x_43, 5, x_33); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_40); -lean_ctor_set(x_43, 8, x_5); -lean_ctor_set_uint8(x_43, sizeof(void*)*9, x_28); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 1, x_29); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 2, x_30); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 3, x_35); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 4, x_36); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 5, x_37); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 6, x_38); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 7, x_39); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 8, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 9, x_42); -x_44 = lean_box(0); -x_45 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_43, x_7, x_8, x_9, x_44, x_11, x_12, x_13); -return x_45; +lean_dec(x_10); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_9); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_45, x_11, x_12, x_13, x_46, x_15, x_16, x_17); +return x_47; } } else { -lean_object* x_46; uint8_t x_47; +lean_object* x_48; uint8_t x_49; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -19037,188 +22174,192 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_46 = lean_ctor_get(x_10, 0); -lean_inc(x_46); -lean_dec(x_10); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +lean_dec(x_14); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_46, 1); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_st_ref_set(x_12, x_49, x_13); -lean_dec(x_12); -x_51 = lean_ctor_get(x_50, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; lean_object* x_56; uint8_t x_57; +x_50 = lean_ctor_get(x_48, 1); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_ctor_get(x_48, 1); -lean_inc(x_52); -x_53 = lean_uint64_of_nat(x_52); +x_52 = lean_st_ref_set(x_16, x_51, x_17); +lean_dec(x_16); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); lean_dec(x_52); -x_54 = lean_io_add_heartbeats(x_53, x_51); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +x_54 = lean_ctor_get(x_50, 1); +lean_inc(x_54); +x_55 = lean_uint64_of_nat(x_54); +lean_dec(x_54); +x_56 = lean_io_add_heartbeats(x_55, x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -lean_ctor_set(x_54, 0, x_46); -return x_54; +lean_object* x_58; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_46); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_48); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint64_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_59 = lean_ctor_get(x_46, 0); -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_46); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint64_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_61 = lean_ctor_get(x_48, 0); +x_62 = lean_ctor_get(x_48, 1); +lean_inc(x_62); lean_inc(x_61); -x_62 = lean_st_ref_set(x_12, x_61, x_13); -lean_dec(x_12); -x_63 = lean_ctor_get(x_62, 1); +lean_dec(x_48); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_62); -x_64 = lean_ctor_get(x_60, 1); -lean_inc(x_64); -x_65 = lean_uint64_of_nat(x_64); +x_64 = lean_st_ref_set(x_16, x_63, x_17); +lean_dec(x_16); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); lean_dec(x_64); -x_66 = lean_io_add_heartbeats(x_65, x_63); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; +x_66 = lean_ctor_get(x_62, 1); +lean_inc(x_66); +x_67 = lean_uint64_of_nat(x_66); +lean_dec(x_66); +x_68 = lean_io_add_heartbeats(x_67, x_65); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; } else { - lean_dec_ref(x_66); - x_68 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_59); -lean_ctor_set(x_69, 1, x_60); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_61); +lean_ctor_set(x_71, 1, x_62); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_70 = x_68; + x_72 = x_70; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 6); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_3, 2); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2), 10, 3); -lean_closure_set(x_18, 0, x_4); -lean_closure_set(x_18, 1, x_3); -lean_closure_set(x_18, 2, x_2); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); -lean_closure_set(x_19, 0, x_17); -lean_closure_set(x_19, 1, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); -lean_closure_set(x_20, 0, x_16); -lean_closure_set(x_20, 1, x_19); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_ctor_get(x_5, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 2); +lean_inc(x_19); +lean_inc(x_7); +x_20 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed), 15, 8); +lean_closure_set(x_20, 0, x_8); +lean_closure_set(x_20, 1, x_6); +lean_closure_set(x_20, 2, x_4); +lean_closure_set(x_20, 3, x_5); +lean_closure_set(x_20, 4, x_7); +lean_closure_set(x_20, 5, x_3); +lean_closure_set(x_20, 6, x_1); +lean_closure_set(x_20, 7, x_2); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); +lean_closure_set(x_21, 0, x_19); +lean_closure_set(x_21, 1, x_20); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); +lean_closure_set(x_22, 0, x_7); +lean_closure_set(x_22, 1, x_21); +x_23 = !lean_is_exclusive(x_10); +if (x_23 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 8); -lean_dec(x_22); -lean_ctor_set(x_6, 8, x_5); -x_23 = lean_box(0); -x_24 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_6, x_7, x_8, x_9, x_23, x_11, x_12, x_13); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_10, 8); +lean_dec(x_24); +lean_ctor_set(x_10, 8, x_9); +x_25 = lean_box(0); +x_26 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_10, x_11, x_12, x_13, x_25, x_15, x_16, x_17); +return x_26; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_6, 0); -x_26 = lean_ctor_get(x_6, 1); -x_27 = lean_ctor_get(x_6, 2); -x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); -x_29 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); -x_30 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); -x_31 = lean_ctor_get(x_6, 3); -x_32 = lean_ctor_get(x_6, 4); -x_33 = lean_ctor_get(x_6, 5); -x_34 = lean_ctor_get(x_6, 6); -x_35 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); -x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); -x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); -x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); -x_39 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); -x_40 = lean_ctor_get(x_6, 7); -x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); -x_42 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); -lean_inc(x_40); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_27 = lean_ctor_get(x_10, 0); +x_28 = lean_ctor_get(x_10, 1); +x_29 = lean_ctor_get(x_10, 2); +x_30 = lean_ctor_get_uint8(x_10, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_10, 3); +x_34 = lean_ctor_get(x_10, 4); +x_35 = lean_ctor_get(x_10, 5); +x_36 = lean_ctor_get(x_10, 6); +x_37 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_10, 7); +x_43 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 9); +lean_inc(x_42); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); +lean_inc(x_29); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_6); -x_43 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_43, 0, x_25); -lean_ctor_set(x_43, 1, x_26); -lean_ctor_set(x_43, 2, x_27); -lean_ctor_set(x_43, 3, x_31); -lean_ctor_set(x_43, 4, x_32); -lean_ctor_set(x_43, 5, x_33); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_40); -lean_ctor_set(x_43, 8, x_5); -lean_ctor_set_uint8(x_43, sizeof(void*)*9, x_28); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 1, x_29); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 2, x_30); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 3, x_35); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 4, x_36); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 5, x_37); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 6, x_38); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 7, x_39); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 8, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 9, x_42); -x_44 = lean_box(0); -x_45 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_43, x_7, x_8, x_9, x_44, x_11, x_12, x_13); -return x_45; +lean_dec(x_10); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_9); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_45, x_11, x_12, x_13, x_46, x_15, x_16, x_17); +return x_47; } } else { -lean_object* x_46; uint8_t x_47; +lean_object* x_48; uint8_t x_49; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -19228,674 +22369,678 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_46 = lean_ctor_get(x_10, 0); -lean_inc(x_46); -lean_dec(x_10); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +lean_dec(x_14); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_46, 1); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_st_ref_set(x_12, x_49, x_13); -lean_dec(x_12); -x_51 = lean_ctor_get(x_50, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; lean_object* x_56; uint8_t x_57; +x_50 = lean_ctor_get(x_48, 1); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_ctor_get(x_48, 1); -lean_inc(x_52); -x_53 = lean_uint64_of_nat(x_52); +x_52 = lean_st_ref_set(x_16, x_51, x_17); +lean_dec(x_16); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); lean_dec(x_52); -x_54 = lean_io_add_heartbeats(x_53, x_51); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +x_54 = lean_ctor_get(x_50, 1); +lean_inc(x_54); +x_55 = lean_uint64_of_nat(x_54); +lean_dec(x_54); +x_56 = lean_io_add_heartbeats(x_55, x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -lean_ctor_set(x_54, 0, x_46); -return x_54; +lean_object* x_58; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_46); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_48); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint64_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_59 = lean_ctor_get(x_46, 0); -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_46); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint64_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_61 = lean_ctor_get(x_48, 0); +x_62 = lean_ctor_get(x_48, 1); +lean_inc(x_62); lean_inc(x_61); -x_62 = lean_st_ref_set(x_12, x_61, x_13); -lean_dec(x_12); -x_63 = lean_ctor_get(x_62, 1); +lean_dec(x_48); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_62); -x_64 = lean_ctor_get(x_60, 1); -lean_inc(x_64); -x_65 = lean_uint64_of_nat(x_64); +x_64 = lean_st_ref_set(x_16, x_63, x_17); +lean_dec(x_16); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); lean_dec(x_64); -x_66 = lean_io_add_heartbeats(x_65, x_63); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; +x_66 = lean_ctor_get(x_62, 1); +lean_inc(x_66); +x_67 = lean_uint64_of_nat(x_66); +lean_dec(x_66); +x_68 = lean_io_add_heartbeats(x_67, x_65); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; } else { - lean_dec_ref(x_66); - x_68 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_59); -lean_ctor_set(x_69, 1, x_60); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_61); +lean_ctor_set(x_71, 1, x_62); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_70 = x_68; + x_72 = x_70; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_box(0); -lean_inc(x_11); -x_16 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_15, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +lean_inc(x_15); +x_20 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_19, x_16, x_17, x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 1); -x_21 = lean_st_ref_get(x_11, x_18); -lean_dec(x_11); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_st_ref_get(x_15, x_22); +lean_dec(x_15); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_17, 1, x_24); -lean_ctor_set(x_21, 0, x_17); -return x_21; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_21, 1, x_28); +lean_ctor_set(x_25, 0, x_21); +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_21, 0); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_21); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_17, 1, x_27); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_26); -return x_28; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_25, 0); +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_25); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_29); +lean_ctor_set(x_21, 1, x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_30); +return x_32; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_17, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_17); -x_31 = lean_st_ref_get(x_11, x_18); -lean_dec(x_11); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 1); +lean_inc(x_34); lean_inc(x_33); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_34 = x_31; +lean_dec(x_21); +x_35 = lean_st_ref_get(x_15, x_22); +lean_dec(x_15); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_38 = x_35; } else { - lean_dec_ref(x_31); - x_34 = lean_box(0); -} -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_32); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_29); -lean_ctor_set(x_36, 1, x_35); -if (lean_is_scalar(x_34)) { - x_37 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_35); + x_38 = lean_box(0); +} +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_39); +if (lean_is_scalar(x_38)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_34; + x_41 = x_38; } -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_33); -return x_37; +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_37); +return x_41; } } else { -uint8_t x_38; -lean_dec(x_11); -x_38 = !lean_is_exclusive(x_16); -if (x_38 == 0) +uint8_t x_42; +lean_dec(x_15); +x_42 = !lean_is_exclusive(x_20); +if (x_42 == 0) { -return x_16; +return x_20; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_16, 0); -x_40 = lean_ctor_get(x_16, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_16); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_20, 0); +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_20); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_1); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = lean_ctor_get(x_1, 0); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_1); +if (x_46 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec(x_45); -lean_ctor_set(x_43, 1, x_46); -lean_inc(x_11); -x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_1, x_12, x_13, x_14); -if (lean_obj_tag(x_47) == 0) +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_1, 0); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); +lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = !lean_is_exclusive(x_48); -if (x_50 == 0) +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +lean_dec(x_49); +lean_ctor_set(x_47, 1, x_50); +lean_inc(x_15); +x_51 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_1, x_16, x_17, x_18); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_48, 1); -x_52 = lean_st_ref_get(x_11, x_49); -lean_dec(x_11); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = !lean_is_exclusive(x_52); +if (x_54 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_51); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_48, 1, x_55); -lean_ctor_set(x_52, 0, x_48); -return x_52; -} -else +lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_55 = lean_ctor_get(x_52, 1); +x_56 = lean_st_ref_get(x_15, x_53); +lean_dec(x_15); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_52, 0); -x_57 = lean_ctor_get(x_52, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_52); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_51); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_48, 1, x_58); +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_57); -return x_59; -} +lean_ctor_set(x_59, 0, x_55); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_52, 1, x_59); +lean_ctor_set(x_56, 0, x_52); +return x_56; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_48); -x_62 = lean_st_ref_get(x_11, x_49); -lean_dec(x_11); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_dec(x_56); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_52, 1, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_52); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_64 = lean_ctor_get(x_52, 0); +x_65 = lean_ctor_get(x_52, 1); +lean_inc(x_65); lean_inc(x_64); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_dec(x_52); +x_66 = lean_st_ref_get(x_15, x_53); +lean_dec(x_15); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_69 = x_66; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_66); + x_69 = lean_box(0); } -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_63); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_60); -lean_ctor_set(x_67, 1, x_66); -if (lean_is_scalar(x_65)) { - x_68 = lean_alloc_ctor(0, 2, 0); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_65); +lean_ctor_set(x_70, 1, x_67); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_64); +lean_ctor_set(x_71, 1, x_70); +if (lean_is_scalar(x_69)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_65; + x_72 = x_69; } -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_64); -return x_68; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_68); +return x_72; } } else { -uint8_t x_69; -lean_dec(x_11); -x_69 = !lean_is_exclusive(x_47); -if (x_69 == 0) +uint8_t x_73; +lean_dec(x_15); +x_73 = !lean_is_exclusive(x_51); +if (x_73 == 0) { -return x_47; +return x_51; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_47, 0); -x_71 = lean_ctor_get(x_47, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_47); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_51, 0); +x_75 = lean_ctor_get(x_51, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_51); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_73 = lean_ctor_get(x_43, 0); -x_74 = lean_ctor_get(x_43, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_43); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -lean_dec(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_73); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_1, 0, x_76); -lean_inc(x_11); -x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_1, x_12, x_13, x_14); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_77, 0); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_47, 0); +x_78 = lean_ctor_get(x_47, 1); lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_77); +lean_dec(x_47); +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_78, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_82 = x_78; -} else { - lean_dec_ref(x_78); - x_82 = lean_box(0); -} -x_83 = lean_st_ref_get(x_11, x_79); -lean_dec(x_11); -x_84 = lean_ctor_get(x_83, 0); +lean_dec(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_1, 0, x_80); +lean_inc(x_15); +x_81 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_1, x_16, x_17, x_18); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); +x_85 = lean_ctor_get(x_82, 1); lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_83); + lean_dec_ref(x_82); x_86 = lean_box(0); } -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_81); -lean_ctor_set(x_87, 1, x_84); -if (lean_is_scalar(x_82)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_87 = lean_st_ref_get(x_15, x_83); +lean_dec(x_15); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; } else { - x_88 = x_82; + lean_dec_ref(x_87); + x_90 = lean_box(0); } -lean_ctor_set(x_88, 0, x_80); -lean_ctor_set(x_88, 1, x_87); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_88); if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(0, 2, 0); + x_92 = lean_alloc_ctor(0, 2, 0); } else { - x_89 = x_86; + x_92 = x_86; } -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_85); -return x_89; +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_91); +if (lean_is_scalar(x_90)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_90; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_89); +return x_93; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_11); -x_90 = lean_ctor_get(x_77, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_77, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_92 = x_77; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_15); +x_94 = lean_ctor_get(x_81, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_81, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_96 = x_81; } else { - lean_dec_ref(x_77); - x_92 = lean_box(0); + lean_dec_ref(x_81); + x_96 = lean_box(0); } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(1, 2, 0); } else { - x_93 = x_92; + x_97 = x_96; } -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_91); -return x_93; +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_95); +return x_97; } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_94 = lean_ctor_get(x_1, 0); -lean_inc(x_94); -lean_dec(x_1); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_97 = x_94; -} else { - lean_dec_ref(x_94); - x_97 = lean_box(0); -} -x_98 = lean_ctor_get(x_96, 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; +x_98 = lean_ctor_get(x_1, 0); lean_inc(x_98); -lean_dec(x_96); -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_1); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_101 = x_98; } else { - x_99 = x_97; + lean_dec_ref(x_98); + x_101 = lean_box(0); } -lean_ctor_set(x_99, 0, x_95); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -lean_inc(x_11); -x_101 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_100, x_12, x_13, x_14); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_102 = lean_ctor_get(x_101, 0); +x_102 = lean_ctor_get(x_100, 0); lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_106 = x_102; +lean_dec(x_100); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_102); - x_106 = lean_box(0); + x_103 = x_101; } -x_107 = lean_st_ref_get(x_11, x_103); -lean_dec(x_11); -x_108 = lean_ctor_get(x_107, 0); +lean_ctor_set(x_103, 0, x_99); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +lean_inc(x_15); +x_105 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_104, x_16, x_17, x_18); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_ctor_get(x_106, 0); lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); +x_109 = lean_ctor_get(x_106, 1); lean_inc(x_109); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_110 = x_107; +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_110 = x_106; } else { - lean_dec_ref(x_107); + lean_dec_ref(x_106); x_110 = lean_box(0); } -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_108); -if (lean_is_scalar(x_106)) { - x_112 = lean_alloc_ctor(0, 2, 0); +x_111 = lean_st_ref_get(x_15, x_107); +lean_dec(x_15); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - x_112 = x_106; + lean_dec_ref(x_111); + x_114 = lean_box(0); } -lean_ctor_set(x_112, 0, x_104); -lean_ctor_set(x_112, 1, x_111); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); if (lean_is_scalar(x_110)) { - x_113 = lean_alloc_ctor(0, 2, 0); + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_110; + x_116 = x_110; } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_109); -return x_113; +lean_ctor_set(x_116, 0, x_108); +lean_ctor_set(x_116, 1, x_115); +if (lean_is_scalar(x_114)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_114; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_113); +return x_117; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_11); -x_114 = lean_ctor_get(x_101, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_101, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_116 = x_101; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_15); +x_118 = lean_ctor_get(x_105, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_105, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_120 = x_105; } else { - lean_dec_ref(x_101); - x_116 = lean_box(0); + lean_dec_ref(x_105); + x_120 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_121 = x_120; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_12) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -x_15 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1(x_8, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14, x_9, x_10, x_11, x_12, x_13); -return x_15; +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1(x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18, x_13, x_14, x_15, x_16, x_17); +return x_19; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_8, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_set(x_10, x_18, x_13); -x_20 = lean_ctor_get(x_19, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_12, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1(x_8, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_20, x_9, x_10, x_11, x_12, x_21); lean_dec(x_20); -return x_22; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_set(x_14, x_22, x_17); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1(x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24, x_13, x_14, x_15, x_16, x_25); +lean_dec(x_24); +return x_26; } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 6); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_3, 2); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2), 10, 3); -lean_closure_set(x_18, 0, x_4); -lean_closure_set(x_18, 1, x_3); -lean_closure_set(x_18, 2, x_2); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); -lean_closure_set(x_19, 0, x_17); -lean_closure_set(x_19, 1, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); -lean_closure_set(x_20, 0, x_16); -lean_closure_set(x_20, 1, x_19); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_ctor_get(x_5, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 2); +lean_inc(x_19); +lean_inc(x_7); +x_20 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed), 15, 8); +lean_closure_set(x_20, 0, x_8); +lean_closure_set(x_20, 1, x_6); +lean_closure_set(x_20, 2, x_4); +lean_closure_set(x_20, 3, x_5); +lean_closure_set(x_20, 4, x_7); +lean_closure_set(x_20, 5, x_3); +lean_closure_set(x_20, 6, x_1); +lean_closure_set(x_20, 7, x_2); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); +lean_closure_set(x_21, 0, x_19); +lean_closure_set(x_21, 1, x_20); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); +lean_closure_set(x_22, 0, x_7); +lean_closure_set(x_22, 1, x_21); +x_23 = !lean_is_exclusive(x_10); +if (x_23 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 8); -lean_dec(x_22); -lean_ctor_set(x_6, 8, x_5); -x_23 = lean_box(0); -x_24 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_6, x_7, x_8, x_9, x_23, x_11, x_12, x_13); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_10, 8); +lean_dec(x_24); +lean_ctor_set(x_10, 8, x_9); +x_25 = lean_box(0); +x_26 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_10, x_11, x_12, x_13, x_25, x_15, x_16, x_17); +return x_26; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_6, 0); -x_26 = lean_ctor_get(x_6, 1); -x_27 = lean_ctor_get(x_6, 2); -x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); -x_29 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); -x_30 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); -x_31 = lean_ctor_get(x_6, 3); -x_32 = lean_ctor_get(x_6, 4); -x_33 = lean_ctor_get(x_6, 5); -x_34 = lean_ctor_get(x_6, 6); -x_35 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); -x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); -x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); -x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); -x_39 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); -x_40 = lean_ctor_get(x_6, 7); -x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); -x_42 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); -lean_inc(x_40); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_27 = lean_ctor_get(x_10, 0); +x_28 = lean_ctor_get(x_10, 1); +x_29 = lean_ctor_get(x_10, 2); +x_30 = lean_ctor_get_uint8(x_10, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_10, 3); +x_34 = lean_ctor_get(x_10, 4); +x_35 = lean_ctor_get(x_10, 5); +x_36 = lean_ctor_get(x_10, 6); +x_37 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_10, 7); +x_43 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 9); +lean_inc(x_42); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); +lean_inc(x_29); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_6); -x_43 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_43, 0, x_25); -lean_ctor_set(x_43, 1, x_26); -lean_ctor_set(x_43, 2, x_27); -lean_ctor_set(x_43, 3, x_31); -lean_ctor_set(x_43, 4, x_32); -lean_ctor_set(x_43, 5, x_33); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_40); -lean_ctor_set(x_43, 8, x_5); -lean_ctor_set_uint8(x_43, sizeof(void*)*9, x_28); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 1, x_29); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 2, x_30); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 3, x_35); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 4, x_36); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 5, x_37); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 6, x_38); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 7, x_39); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 8, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 9, x_42); -x_44 = lean_box(0); -x_45 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_43, x_7, x_8, x_9, x_44, x_11, x_12, x_13); -return x_45; +lean_dec(x_10); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_9); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_45, x_11, x_12, x_13, x_46, x_15, x_16, x_17); +return x_47; } } else { -lean_object* x_46; uint8_t x_47; +lean_object* x_48; uint8_t x_49; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -19904,189 +23049,193 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_46 = lean_ctor_get(x_10, 0); -lean_inc(x_46); -lean_dec(x_10); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +lean_dec(x_1); +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +lean_dec(x_14); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_46, 1); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_st_ref_set(x_12, x_49, x_13); -lean_dec(x_12); -x_51 = lean_ctor_get(x_50, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; lean_object* x_56; uint8_t x_57; +x_50 = lean_ctor_get(x_48, 1); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_ctor_get(x_48, 1); -lean_inc(x_52); -x_53 = lean_uint64_of_nat(x_52); +x_52 = lean_st_ref_set(x_16, x_51, x_17); +lean_dec(x_16); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); lean_dec(x_52); -x_54 = lean_io_add_heartbeats(x_53, x_51); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +x_54 = lean_ctor_get(x_50, 1); +lean_inc(x_54); +x_55 = lean_uint64_of_nat(x_54); +lean_dec(x_54); +x_56 = lean_io_add_heartbeats(x_55, x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -lean_ctor_set(x_54, 0, x_46); -return x_54; +lean_object* x_58; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_46); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_48); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint64_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_59 = lean_ctor_get(x_46, 0); -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_46); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint64_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_61 = lean_ctor_get(x_48, 0); +x_62 = lean_ctor_get(x_48, 1); +lean_inc(x_62); lean_inc(x_61); -x_62 = lean_st_ref_set(x_12, x_61, x_13); -lean_dec(x_12); -x_63 = lean_ctor_get(x_62, 1); +lean_dec(x_48); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_62); -x_64 = lean_ctor_get(x_60, 1); -lean_inc(x_64); -x_65 = lean_uint64_of_nat(x_64); +x_64 = lean_st_ref_set(x_16, x_63, x_17); +lean_dec(x_16); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); lean_dec(x_64); -x_66 = lean_io_add_heartbeats(x_65, x_63); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; +x_66 = lean_ctor_get(x_62, 1); +lean_inc(x_66); +x_67 = lean_uint64_of_nat(x_66); +lean_dec(x_66); +x_68 = lean_io_add_heartbeats(x_67, x_65); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; } else { - lean_dec_ref(x_66); - x_68 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_59); -lean_ctor_set(x_69, 1, x_60); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_61); +lean_ctor_set(x_71, 1, x_62); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_70 = x_68; + x_72 = x_70; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } } } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 6); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_3, 2); -lean_inc(x_17); -x_18 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__2), 10, 3); -lean_closure_set(x_18, 0, x_4); -lean_closure_set(x_18, 1, x_3); -lean_closure_set(x_18, 2, x_2); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); -lean_closure_set(x_19, 0, x_17); -lean_closure_set(x_19, 1, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); -lean_closure_set(x_20, 0, x_16); -lean_closure_set(x_20, 1, x_19); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_ctor_get(x_5, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 2); +lean_inc(x_19); +lean_inc(x_7); +x_20 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed), 15, 8); +lean_closure_set(x_20, 0, x_8); +lean_closure_set(x_20, 1, x_6); +lean_closure_set(x_20, 2, x_4); +lean_closure_set(x_20, 3, x_5); +lean_closure_set(x_20, 4, x_7); +lean_closure_set(x_20, 5, x_3); +lean_closure_set(x_20, 6, x_1); +lean_closure_set(x_20, 7, x_2); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withLevelNames___rarg), 9, 2); +lean_closure_set(x_21, 0, x_19); +lean_closure_set(x_21, 1, x_20); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withDeclName___rarg), 9, 2); +lean_closure_set(x_22, 0, x_7); +lean_closure_set(x_22, 1, x_21); +x_23 = !lean_is_exclusive(x_10); +if (x_23 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 8); -lean_dec(x_22); -lean_ctor_set(x_6, 8, x_5); -x_23 = lean_box(0); -x_24 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_6, x_7, x_8, x_9, x_23, x_11, x_12, x_13); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_10, 8); +lean_dec(x_24); +lean_ctor_set(x_10, 8, x_9); +x_25 = lean_box(0); +x_26 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_10, x_11, x_12, x_13, x_25, x_15, x_16, x_17); +return x_26; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_6, 0); -x_26 = lean_ctor_get(x_6, 1); -x_27 = lean_ctor_get(x_6, 2); -x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); -x_29 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 1); -x_30 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 2); -x_31 = lean_ctor_get(x_6, 3); -x_32 = lean_ctor_get(x_6, 4); -x_33 = lean_ctor_get(x_6, 5); -x_34 = lean_ctor_get(x_6, 6); -x_35 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 3); -x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 4); -x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 5); -x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 6); -x_39 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 7); -x_40 = lean_ctor_get(x_6, 7); -x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 8); -x_42 = lean_ctor_get_uint8(x_6, sizeof(void*)*9 + 9); -lean_inc(x_40); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_27 = lean_ctor_get(x_10, 0); +x_28 = lean_ctor_get(x_10, 1); +x_29 = lean_ctor_get(x_10, 2); +x_30 = lean_ctor_get_uint8(x_10, sizeof(void*)*9); +x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 1); +x_32 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 2); +x_33 = lean_ctor_get(x_10, 3); +x_34 = lean_ctor_get(x_10, 4); +x_35 = lean_ctor_get(x_10, 5); +x_36 = lean_ctor_get(x_10, 6); +x_37 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 3); +x_38 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 4); +x_39 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 5); +x_40 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 6); +x_41 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 7); +x_42 = lean_ctor_get(x_10, 7); +x_43 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 8); +x_44 = lean_ctor_get_uint8(x_10, sizeof(void*)*9 + 9); +lean_inc(x_42); +lean_inc(x_36); +lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); +lean_inc(x_29); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_6); -x_43 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_43, 0, x_25); -lean_ctor_set(x_43, 1, x_26); -lean_ctor_set(x_43, 2, x_27); -lean_ctor_set(x_43, 3, x_31); -lean_ctor_set(x_43, 4, x_32); -lean_ctor_set(x_43, 5, x_33); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_40); -lean_ctor_set(x_43, 8, x_5); -lean_ctor_set_uint8(x_43, sizeof(void*)*9, x_28); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 1, x_29); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 2, x_30); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 3, x_35); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 4, x_36); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 5, x_37); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 6, x_38); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 7, x_39); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 8, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*9 + 9, x_42); -x_44 = lean_box(0); -x_45 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_15, x_20, x_43, x_7, x_8, x_9, x_44, x_11, x_12, x_13); -return x_45; +lean_dec(x_10); +x_45 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_45, 0, x_27); +lean_ctor_set(x_45, 1, x_28); +lean_ctor_set(x_45, 2, x_29); +lean_ctor_set(x_45, 3, x_33); +lean_ctor_set(x_45, 4, x_34); +lean_ctor_set(x_45, 5, x_35); +lean_ctor_set(x_45, 6, x_36); +lean_ctor_set(x_45, 7, x_42); +lean_ctor_set(x_45, 8, x_9); +lean_ctor_set_uint8(x_45, sizeof(void*)*9, x_30); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 1, x_31); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 2, x_32); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 3, x_37); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 4, x_38); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 5, x_39); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 6, x_40); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 7, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 8, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*9 + 9, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_18, x_22, x_45, x_11, x_12, x_13, x_46, x_15, x_16, x_17); +return x_47; } } else { -lean_object* x_46; uint8_t x_47; +lean_object* x_48; uint8_t x_49; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -20096,1152 +23245,1146 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_46 = lean_ctor_get(x_10, 0); -lean_inc(x_46); -lean_dec(x_10); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +lean_dec(x_14); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_46, 1); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_st_ref_set(x_12, x_49, x_13); -lean_dec(x_12); -x_51 = lean_ctor_get(x_50, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; lean_object* x_56; uint8_t x_57; +x_50 = lean_ctor_get(x_48, 1); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_ctor_get(x_48, 1); -lean_inc(x_52); -x_53 = lean_uint64_of_nat(x_52); +x_52 = lean_st_ref_set(x_16, x_51, x_17); +lean_dec(x_16); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); lean_dec(x_52); -x_54 = lean_io_add_heartbeats(x_53, x_51); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +x_54 = lean_ctor_get(x_50, 1); +lean_inc(x_54); +x_55 = lean_uint64_of_nat(x_54); +lean_dec(x_54); +x_56 = lean_io_add_heartbeats(x_55, x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -lean_ctor_set(x_54, 0, x_46); -return x_54; +lean_object* x_58; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_46); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_48); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint64_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_59 = lean_ctor_get(x_46, 0); -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_46); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint64_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_61 = lean_ctor_get(x_48, 0); +x_62 = lean_ctor_get(x_48, 1); +lean_inc(x_62); lean_inc(x_61); -x_62 = lean_st_ref_set(x_12, x_61, x_13); -lean_dec(x_12); -x_63 = lean_ctor_get(x_62, 1); +lean_dec(x_48); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_62); -x_64 = lean_ctor_get(x_60, 1); -lean_inc(x_64); -x_65 = lean_uint64_of_nat(x_64); +x_64 = lean_st_ref_set(x_16, x_63, x_17); +lean_dec(x_16); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); lean_dec(x_64); -x_66 = lean_io_add_heartbeats(x_65, x_63); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; +x_66 = lean_ctor_get(x_62, 1); +lean_inc(x_66); +x_67 = lean_uint64_of_nat(x_66); +lean_dec(x_66); +x_68 = lean_io_add_heartbeats(x_67, x_65); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; } else { - lean_dec_ref(x_66); - x_68 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_59); -lean_ctor_set(x_69, 1, x_60); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_61); +lean_ctor_set(x_71, 1, x_62); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_70 = x_68; + x_72 = x_70; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_box(0); -lean_inc(x_11); -x_16 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_15, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +lean_inc(x_15); +x_20 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_19, x_16, x_17, x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 1); -x_21 = lean_st_ref_get(x_11, x_18); -lean_dec(x_11); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_st_ref_get(x_15, x_22); +lean_dec(x_15); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_17, 1, x_24); -lean_ctor_set(x_21, 0, x_17); -return x_21; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_21, 1, x_28); +lean_ctor_set(x_25, 0, x_21); +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_21, 0); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_21); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_17, 1, x_27); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_26); -return x_28; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_25, 0); +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_25); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_29); +lean_ctor_set(x_21, 1, x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_30); +return x_32; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_17, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_17); -x_31 = lean_st_ref_get(x_11, x_18); -lean_dec(x_11); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 1); +lean_inc(x_34); lean_inc(x_33); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_34 = x_31; +lean_dec(x_21); +x_35 = lean_st_ref_get(x_15, x_22); +lean_dec(x_15); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_38 = x_35; } else { - lean_dec_ref(x_31); - x_34 = lean_box(0); + lean_dec_ref(x_35); + x_38 = lean_box(0); } -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_32); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_29); -lean_ctor_set(x_36, 1, x_35); -if (lean_is_scalar(x_34)) { - x_37 = lean_alloc_ctor(0, 2, 0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_39); +if (lean_is_scalar(x_38)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_34; + x_41 = x_38; } -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_33); -return x_37; +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_37); +return x_41; } } else { -uint8_t x_38; -lean_dec(x_11); -x_38 = !lean_is_exclusive(x_16); -if (x_38 == 0) +uint8_t x_42; +lean_dec(x_15); +x_42 = !lean_is_exclusive(x_20); +if (x_42 == 0) { -return x_16; +return x_20; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_16, 0); -x_40 = lean_ctor_get(x_16, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_16); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_20, 0); +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_20); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_1); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = lean_ctor_get(x_1, 0); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_1); +if (x_46 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec(x_45); -lean_ctor_set(x_43, 1, x_46); -lean_inc(x_11); -x_47 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_1, x_12, x_13, x_14); -if (lean_obj_tag(x_47) == 0) +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_1, 0); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); +lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = !lean_is_exclusive(x_48); -if (x_50 == 0) +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +lean_dec(x_49); +lean_ctor_set(x_47, 1, x_50); +lean_inc(x_15); +x_51 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_1, x_16, x_17, x_18); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_48, 1); -x_52 = lean_st_ref_get(x_11, x_49); -lean_dec(x_11); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = !lean_is_exclusive(x_52); +if (x_54 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_51); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_48, 1, x_55); -lean_ctor_set(x_52, 0, x_48); -return x_52; -} -else +lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_55 = lean_ctor_get(x_52, 1); +x_56 = lean_st_ref_get(x_15, x_53); +lean_dec(x_15); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_52, 0); -x_57 = lean_ctor_get(x_52, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_52); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_51); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_48, 1, x_58); +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_57); -return x_59; -} +lean_ctor_set(x_59, 0, x_55); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_52, 1, x_59); +lean_ctor_set(x_56, 0, x_52); +return x_56; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_48); -x_62 = lean_st_ref_get(x_11, x_49); -lean_dec(x_11); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_dec(x_56); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_52, 1, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_52); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_64 = lean_ctor_get(x_52, 0); +x_65 = lean_ctor_get(x_52, 1); +lean_inc(x_65); lean_inc(x_64); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_dec(x_52); +x_66 = lean_st_ref_get(x_15, x_53); +lean_dec(x_15); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_69 = x_66; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_66); + x_69 = lean_box(0); } -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_63); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_60); -lean_ctor_set(x_67, 1, x_66); -if (lean_is_scalar(x_65)) { - x_68 = lean_alloc_ctor(0, 2, 0); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_65); +lean_ctor_set(x_70, 1, x_67); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_64); +lean_ctor_set(x_71, 1, x_70); +if (lean_is_scalar(x_69)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_65; + x_72 = x_69; } -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_64); -return x_68; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_68); +return x_72; } } else { -uint8_t x_69; -lean_dec(x_11); -x_69 = !lean_is_exclusive(x_47); -if (x_69 == 0) +uint8_t x_73; +lean_dec(x_15); +x_73 = !lean_is_exclusive(x_51); +if (x_73 == 0) { -return x_47; +return x_51; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_47, 0); -x_71 = lean_ctor_get(x_47, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_47); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_51, 0); +x_75 = lean_ctor_get(x_51, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_51); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_73 = lean_ctor_get(x_43, 0); -x_74 = lean_ctor_get(x_43, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_43); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -lean_dec(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_73); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_1, 0, x_76); -lean_inc(x_11); -x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_1, x_12, x_13, x_14); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_77, 0); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_47, 0); +x_78 = lean_ctor_get(x_47, 1); lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_77); +lean_dec(x_47); +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_78, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_82 = x_78; -} else { - lean_dec_ref(x_78); - x_82 = lean_box(0); -} -x_83 = lean_st_ref_get(x_11, x_79); -lean_dec(x_11); -x_84 = lean_ctor_get(x_83, 0); +lean_dec(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_1, 0, x_80); +lean_inc(x_15); +x_81 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_1, x_16, x_17, x_18); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); +x_85 = lean_ctor_get(x_82, 1); lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_83); + lean_dec_ref(x_82); x_86 = lean_box(0); } -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_81); -lean_ctor_set(x_87, 1, x_84); -if (lean_is_scalar(x_82)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_87 = lean_st_ref_get(x_15, x_83); +lean_dec(x_15); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; } else { - x_88 = x_82; + lean_dec_ref(x_87); + x_90 = lean_box(0); } -lean_ctor_set(x_88, 0, x_80); -lean_ctor_set(x_88, 1, x_87); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_88); if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(0, 2, 0); + x_92 = lean_alloc_ctor(0, 2, 0); } else { - x_89 = x_86; + x_92 = x_86; } -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_85); -return x_89; +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_91); +if (lean_is_scalar(x_90)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_90; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_89); +return x_93; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_11); -x_90 = lean_ctor_get(x_77, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_77, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_92 = x_77; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_15); +x_94 = lean_ctor_get(x_81, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_81, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_96 = x_81; } else { - lean_dec_ref(x_77); - x_92 = lean_box(0); + lean_dec_ref(x_81); + x_96 = lean_box(0); } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(1, 2, 0); } else { - x_93 = x_92; + x_97 = x_96; } -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_91); -return x_93; +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_95); +return x_97; } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_94 = lean_ctor_get(x_1, 0); -lean_inc(x_94); -lean_dec(x_1); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_97 = x_94; -} else { - lean_dec_ref(x_94); - x_97 = lean_box(0); -} -x_98 = lean_ctor_get(x_96, 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; +x_98 = lean_ctor_get(x_1, 0); lean_inc(x_98); -lean_dec(x_96); -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_1); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_101 = x_98; } else { - x_99 = x_97; + lean_dec_ref(x_98); + x_101 = lean_box(0); } -lean_ctor_set(x_99, 0, x_95); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -lean_inc(x_11); -x_101 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_100, x_12, x_13, x_14); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_102 = lean_ctor_get(x_101, 0); +x_102 = lean_ctor_get(x_100, 0); lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_106 = x_102; +lean_dec(x_100); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_102); - x_106 = lean_box(0); + x_103 = x_101; } -x_107 = lean_st_ref_get(x_11, x_103); -lean_dec(x_11); -x_108 = lean_ctor_get(x_107, 0); +lean_ctor_set(x_103, 0, x_99); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +lean_inc(x_15); +x_105 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_14, x_15, x_104, x_16, x_17, x_18); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_ctor_get(x_106, 0); lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); +x_109 = lean_ctor_get(x_106, 1); lean_inc(x_109); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_110 = x_107; +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_110 = x_106; } else { - lean_dec_ref(x_107); + lean_dec_ref(x_106); x_110 = lean_box(0); } -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_108); -if (lean_is_scalar(x_106)) { - x_112 = lean_alloc_ctor(0, 2, 0); +x_111 = lean_st_ref_get(x_15, x_107); +lean_dec(x_15); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - x_112 = x_106; + lean_dec_ref(x_111); + x_114 = lean_box(0); } -lean_ctor_set(x_112, 0, x_104); -lean_ctor_set(x_112, 1, x_111); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); if (lean_is_scalar(x_110)) { - x_113 = lean_alloc_ctor(0, 2, 0); + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_110; + x_116 = x_110; } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_109); -return x_113; +lean_ctor_set(x_116, 0, x_108); +lean_ctor_set(x_116, 1, x_115); +if (lean_is_scalar(x_114)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_114; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_113); +return x_117; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_11); -x_114 = lean_ctor_get(x_101, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_101, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_116 = x_101; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_15); +x_118 = lean_ctor_get(x_105, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_105, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_120 = x_105; } else { - lean_dec_ref(x_101); - x_116 = lean_box(0); + lean_dec_ref(x_105); + x_120 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_121 = x_120; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_12) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -x_15 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1(x_8, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14, x_9, x_10, x_11, x_12, x_13); -return x_15; +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1(x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18, x_13, x_14, x_15, x_16, x_17); +return x_19; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_8, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_set(x_10, x_18, x_13); -x_20 = lean_ctor_get(x_19, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_12, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1(x_8, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_20, x_9, x_10, x_11, x_12, x_21); lean_dec(x_20); -return x_22; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_set(x_14, x_22, x_17); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1(x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24, x_13, x_14, x_15, x_16, x_25); +lean_dec(x_24); +return x_26; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_box(0); -lean_inc(x_9); -x_16 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_15, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +lean_inc(x_13); +x_20 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12, x_13, x_19, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 1); -x_21 = lean_st_ref_get(x_9, x_18); -lean_dec(x_9); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_st_ref_get(x_13, x_22); +lean_dec(x_13); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_17, 1, x_24); -lean_ctor_set(x_21, 0, x_17); -return x_21; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_21, 1, x_28); +lean_ctor_set(x_25, 0, x_21); +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_21, 0); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_21); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_17, 1, x_27); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_26); -return x_28; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_25, 0); +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_25); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_29); +lean_ctor_set(x_21, 1, x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_30); +return x_32; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_17, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_17); -x_31 = lean_st_ref_get(x_9, x_18); -lean_dec(x_9); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 1); +lean_inc(x_34); lean_inc(x_33); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_34 = x_31; +lean_dec(x_21); +x_35 = lean_st_ref_get(x_13, x_22); +lean_dec(x_13); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_38 = x_35; } else { - lean_dec_ref(x_31); - x_34 = lean_box(0); + lean_dec_ref(x_35); + x_38 = lean_box(0); } -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_32); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_29); -lean_ctor_set(x_36, 1, x_35); -if (lean_is_scalar(x_34)) { - x_37 = lean_alloc_ctor(0, 2, 0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_39); +if (lean_is_scalar(x_38)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_34; + x_41 = x_38; } -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_33); -return x_37; +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_37); +return x_41; } } else { -uint8_t x_38; -lean_dec(x_9); -x_38 = !lean_is_exclusive(x_16); -if (x_38 == 0) +uint8_t x_42; +lean_dec(x_13); +x_42 = !lean_is_exclusive(x_20); +if (x_42 == 0) { -return x_16; +return x_20; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_16, 0); -x_40 = lean_ctor_get(x_16, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_16); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_20, 0); +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_20); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_1); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = lean_ctor_get(x_1, 0); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_1); +if (x_46 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_43, 1); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec(x_45); -lean_ctor_set(x_43, 1, x_46); -lean_inc(x_9); -x_47 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_1, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_47) == 0) +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_1, 0); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); +lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = !lean_is_exclusive(x_48); -if (x_50 == 0) +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +lean_dec(x_49); +lean_ctor_set(x_47, 1, x_50); +lean_inc(x_13); +x_51 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12, x_13, x_1, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_48, 1); -x_52 = lean_st_ref_get(x_9, x_49); -lean_dec(x_9); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = !lean_is_exclusive(x_52); +if (x_54 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_51); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_48, 1, x_55); -lean_ctor_set(x_52, 0, x_48); -return x_52; -} -else +lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_55 = lean_ctor_get(x_52, 1); +x_56 = lean_st_ref_get(x_13, x_53); +lean_dec(x_13); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_52, 0); -x_57 = lean_ctor_get(x_52, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_52); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_51); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_48, 1, x_58); +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_57); -return x_59; -} +lean_ctor_set(x_59, 0, x_55); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_52, 1, x_59); +lean_ctor_set(x_56, 0, x_52); +return x_56; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_48); -x_62 = lean_st_ref_get(x_9, x_49); -lean_dec(x_9); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_dec(x_56); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_52, 1, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_52); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_64 = lean_ctor_get(x_52, 0); +x_65 = lean_ctor_get(x_52, 1); +lean_inc(x_65); lean_inc(x_64); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_dec(x_52); +x_66 = lean_st_ref_get(x_13, x_53); +lean_dec(x_13); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_69 = x_66; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_66); + x_69 = lean_box(0); } -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_63); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_60); -lean_ctor_set(x_67, 1, x_66); -if (lean_is_scalar(x_65)) { - x_68 = lean_alloc_ctor(0, 2, 0); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_65); +lean_ctor_set(x_70, 1, x_67); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_64); +lean_ctor_set(x_71, 1, x_70); +if (lean_is_scalar(x_69)) { + x_72 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_65; + x_72 = x_69; } -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_64); -return x_68; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_68); +return x_72; } } else { -uint8_t x_69; -lean_dec(x_9); -x_69 = !lean_is_exclusive(x_47); -if (x_69 == 0) +uint8_t x_73; +lean_dec(x_13); +x_73 = !lean_is_exclusive(x_51); +if (x_73 == 0) { -return x_47; +return x_51; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_47, 0); -x_71 = lean_ctor_get(x_47, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_47); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_51, 0); +x_75 = lean_ctor_get(x_51, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_51); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_73 = lean_ctor_get(x_43, 0); -x_74 = lean_ctor_get(x_43, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_43); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -lean_dec(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_73); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_1, 0, x_76); -lean_inc(x_9); -x_77 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_1, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_77, 0); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_47, 0); +x_78 = lean_ctor_get(x_47, 1); lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_77); +lean_dec(x_47); +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_78, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_82 = x_78; -} else { - lean_dec_ref(x_78); - x_82 = lean_box(0); -} -x_83 = lean_st_ref_get(x_9, x_79); -lean_dec(x_9); -x_84 = lean_ctor_get(x_83, 0); +lean_dec(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_1, 0, x_80); +lean_inc(x_13); +x_81 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12, x_13, x_1, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); +x_85 = lean_ctor_get(x_82, 1); lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_83); + lean_dec_ref(x_82); x_86 = lean_box(0); } -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_81); -lean_ctor_set(x_87, 1, x_84); -if (lean_is_scalar(x_82)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_87 = lean_st_ref_get(x_13, x_83); +lean_dec(x_13); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; } else { - x_88 = x_82; + lean_dec_ref(x_87); + x_90 = lean_box(0); } -lean_ctor_set(x_88, 0, x_80); -lean_ctor_set(x_88, 1, x_87); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_88); if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(0, 2, 0); + x_92 = lean_alloc_ctor(0, 2, 0); } else { - x_89 = x_86; + x_92 = x_86; } -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_85); -return x_89; +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_91); +if (lean_is_scalar(x_90)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_90; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_89); +return x_93; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_9); -x_90 = lean_ctor_get(x_77, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_77, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_92 = x_77; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_13); +x_94 = lean_ctor_get(x_81, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_81, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_96 = x_81; } else { - lean_dec_ref(x_77); - x_92 = lean_box(0); + lean_dec_ref(x_81); + x_96 = lean_box(0); } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(1, 2, 0); } else { - x_93 = x_92; + x_97 = x_96; } -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_91); -return x_93; +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_95); +return x_97; } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_94 = lean_ctor_get(x_1, 0); -lean_inc(x_94); -lean_dec(x_1); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_94)) { - lean_ctor_release(x_94, 0); - lean_ctor_release(x_94, 1); - x_97 = x_94; -} else { - lean_dec_ref(x_94); - x_97 = lean_box(0); -} -x_98 = lean_ctor_get(x_96, 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; +x_98 = lean_ctor_get(x_1, 0); lean_inc(x_98); -lean_dec(x_96); -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_1); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_101 = x_98; } else { - x_99 = x_97; + lean_dec_ref(x_98); + x_101 = lean_box(0); } -lean_ctor_set(x_99, 0, x_95); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -lean_inc(x_9); -x_101 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7(x_2, x_3, x_4, x_5, x_6, x_8, x_9, x_100, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_102 = lean_ctor_get(x_101, 0); +x_102 = lean_ctor_get(x_100, 0); lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_106 = x_102; +lean_dec(x_100); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_102); - x_106 = lean_box(0); + x_103 = x_101; } -x_107 = lean_st_ref_get(x_9, x_103); -lean_dec(x_9); -x_108 = lean_ctor_get(x_107, 0); +lean_ctor_set(x_103, 0, x_99); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +lean_inc(x_13); +x_105 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12, x_13, x_104, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_ctor_get(x_106, 0); lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); +x_109 = lean_ctor_get(x_106, 1); lean_inc(x_109); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_110 = x_107; +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_110 = x_106; } else { - lean_dec_ref(x_107); + lean_dec_ref(x_106); x_110 = lean_box(0); } -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_108); -if (lean_is_scalar(x_106)) { - x_112 = lean_alloc_ctor(0, 2, 0); +x_111 = lean_st_ref_get(x_13, x_107); +lean_dec(x_13); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - x_112 = x_106; + lean_dec_ref(x_111); + x_114 = lean_box(0); } -lean_ctor_set(x_112, 0, x_104); -lean_ctor_set(x_112, 1, x_111); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); if (lean_is_scalar(x_110)) { - x_113 = lean_alloc_ctor(0, 2, 0); + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_110; + x_116 = x_110; } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_109); -return x_113; +lean_ctor_set(x_116, 0, x_108); +lean_ctor_set(x_116, 1, x_115); +if (lean_is_scalar(x_114)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_114; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_113); +return x_117; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_9); -x_114 = lean_ctor_get(x_101, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_101, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_116 = x_101; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_13); +x_118 = lean_ctor_get(x_105, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_105, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_120 = x_105; } else { - lean_dec_ref(x_101); - x_116 = lean_box(0); + lean_dec_ref(x_105); + x_120 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_121 = x_120; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_9) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -x_15 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(x_5, x_1, x_2, x_3, x_4, x_6, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_18, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_19; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_5, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_st_ref_set(x_8, x_18, x_13); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_19, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_9, 0); lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(x_5, x_1, x_2, x_3, x_4, x_6, x_21, x_7, x_8, x_9, x_10, x_11, x_12, x_20); -return x_22; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_st_ref_set(x_12, x_22, x_17); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_25, x_11, x_12, x_13, x_14, x_15, x_16, x_24); +return x_26; } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_6, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_10, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -lean_dec(x_23); +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_dec(x_27); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_ctor_get(x_19, 1); -lean_inc(x_25); -lean_dec(x_19); -x_26 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___closed__2; -x_27 = l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__11(x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_25); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_dec(x_23); +x_30 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___closed__2; +x_31 = l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__11(x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_29); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -return x_27; +return x_31; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_27); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_31); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_32 = lean_ctor_get(x_19, 1); -lean_inc(x_32); -lean_dec(x_19); -x_33 = lean_ctor_get(x_24, 0); -lean_inc(x_33); -lean_dec(x_24); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = l_Lean_Language_SnapshotTask_get___rarg(x_34); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); lean_dec(x_23); -x_37 = lean_io_promise_resolve(x_35, x_36, x_32); -lean_dec(x_36); -x_38 = lean_ctor_get(x_37, 0); +x_37 = lean_ctor_get(x_28, 0); +lean_inc(x_37); +lean_dec(x_28); +x_38 = lean_ctor_get(x_37, 1); lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); lean_dec(x_37); -x_40 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(x_5, x_1, x_2, x_3, x_4, x_6, x_38, x_7, x_8, x_9, x_10, x_11, x_12, x_39); -lean_dec(x_38); -return x_40; -} -} +x_39 = l_Lean_Language_SnapshotTask_get___rarg(x_38); +x_40 = lean_ctor_get(x_27, 1); +lean_inc(x_40); +lean_dec(x_27); +x_41 = lean_io_promise_resolve(x_39, x_40, x_36); +lean_dec(x_40); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_42, x_11, x_12, x_13, x_14, x_15, x_16, x_43); +lean_dec(x_42); +return x_44; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_9); -return x_10; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -21249,883 +24392,918 @@ x_1 = lean_mk_string_unchecked("elabFunValues", 13, 13); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__10; -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_12 = lean_ctor_get(x_1, 2); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 6); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_ctor_get(x_1, 2); lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 0); lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm), 3, 1); -lean_closure_set(x_16, 0, x_14); -lean_inc(x_10); +x_16 = lean_ctor_get(x_15, 6); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm), 3, 1); +lean_closure_set(x_19, 0, x_16); lean_inc(x_12); +lean_inc(x_14); lean_inc(x_1); -x_17 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_1, x_12, x_15, x_16, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_17) == 0) +x_20 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_2, x_3, x_1, x_14, x_15, x_17, x_18, x_19, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 3); -lean_inc(x_19); +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 3); +lean_inc(x_22); lean_dec(x_1); -if (lean_obj_tag(x_19) == 0) +if (lean_obj_tag(x_22) == 0) { -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_2); -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +uint8_t x_23; +lean_dec(x_12); +lean_dec(x_4); +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_17, 0); +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_20, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_21, 0); +lean_inc(x_25); lean_dec(x_21); -x_22 = lean_ctor_get(x_18, 0); -lean_inc(x_22); -lean_dec(x_18); -lean_ctor_set(x_17, 0, x_22); -return x_17; +lean_ctor_set(x_20, 0, x_25); +return x_20; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); -lean_dec(x_17); -x_24 = lean_ctor_get(x_18, 0); -lean_inc(x_24); -lean_dec(x_18); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = lean_ctor_get(x_21, 0); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_17, 1); -lean_inc(x_26); -lean_dec(x_17); -x_27 = lean_ctor_get(x_18, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_18, 1); -lean_inc(x_28); -lean_dec(x_18); -x_29 = !lean_is_exclusive(x_19); -if (x_29 == 0) +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_dec(x_20); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_21, 1); +lean_inc(x_31); +lean_dec(x_21); +x_32 = !lean_is_exclusive(x_22); +if (x_32 == 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; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_30 = lean_ctor_get(x_19, 0); -x_31 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_10, x_26); -lean_dec(x_10); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_32, x_33); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_33 = lean_ctor_get(x_22, 0); +x_34 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_12, x_29); +lean_dec(x_12); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); lean_dec(x_34); -x_37 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3; -x_38 = 0; -x_39 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_35); -lean_ctor_set(x_39, 2, x_2); -lean_ctor_set_uint8(x_39, sizeof(void*)*3, x_38); -lean_inc(x_27); -x_40 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_28); -lean_ctor_set(x_40, 2, x_27); -lean_ctor_set(x_19, 0, x_40); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_dec(x_30); -x_42 = lean_io_promise_resolve(x_19, x_41, x_36); -lean_dec(x_41); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) -{ -lean_object* x_44; -x_44 = lean_ctor_get(x_42, 0); +x_37 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_35, x_36); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3; +x_41 = 0; +x_42 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_38); +lean_ctor_set(x_42, 2, x_4); +lean_ctor_set_uint8(x_42, sizeof(void*)*3, x_41); +lean_inc(x_30); +x_43 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_31); +lean_ctor_set(x_43, 2, x_30); +lean_ctor_set(x_22, 0, x_43); +x_44 = lean_ctor_get(x_33, 1); +lean_inc(x_44); +lean_dec(x_33); +x_45 = lean_io_promise_resolve(x_22, x_44, x_39); lean_dec(x_44); -lean_ctor_set(x_42, 0, x_27); -return x_42; +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +lean_object* x_47; +x_47 = lean_ctor_get(x_45, 0); +lean_dec(x_47); +lean_ctor_set(x_45, 0, x_30); +return x_45; } else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_27); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_30); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_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; -x_47 = lean_ctor_get(x_19, 0); -lean_inc(x_47); -lean_dec(x_19); -x_48 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_10, x_26); -lean_dec(x_10); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; 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; +x_50 = lean_ctor_get(x_22, 0); lean_inc(x_50); -lean_dec(x_48); -x_51 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_49, x_50); +lean_dec(x_22); +x_51 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_12, x_29); +lean_dec(x_12); x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3; -x_55 = 0; -x_56 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_52); -lean_ctor_set(x_56, 2, x_2); -lean_ctor_set_uint8(x_56, sizeof(void*)*3, x_55); -lean_inc(x_27); -x_57 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_28); -lean_ctor_set(x_57, 2, x_27); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_59 = lean_ctor_get(x_47, 1); -lean_inc(x_59); -lean_dec(x_47); -x_60 = lean_io_promise_resolve(x_58, x_59, x_53); -lean_dec(x_59); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_62 = x_60; +x_54 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_52, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3; +x_58 = 0; +x_59 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_55); +lean_ctor_set(x_59, 2, x_4); +lean_ctor_set_uint8(x_59, sizeof(void*)*3, x_58); +lean_inc(x_30); +x_60 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_31); +lean_ctor_set(x_60, 2, x_30); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = lean_ctor_get(x_50, 1); +lean_inc(x_62); +lean_dec(x_50); +x_63 = lean_io_promise_resolve(x_61, x_62, x_56); +lean_dec(x_62); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_65 = x_63; } else { - lean_dec_ref(x_60); - x_62 = lean_box(0); + lean_dec_ref(x_63); + x_65 = lean_box(0); } -if (lean_is_scalar(x_62)) { - x_63 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_62; + x_66 = x_65; } -lean_ctor_set(x_63, 0, x_27); -lean_ctor_set(x_63, 1, x_61); -return x_63; +lean_ctor_set(x_66, 0, x_30); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } } else { -uint8_t x_64; -lean_dec(x_10); -lean_dec(x_2); +uint8_t x_67; +lean_dec(x_12); +lean_dec(x_4); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_17); -if (x_64 == 0) +x_67 = !lean_is_exclusive(x_20); +if (x_67 == 0) { -return x_17; +return x_20; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_17, 0); -x_66 = lean_ctor_get(x_17, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_17); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_20, 0); +x_69 = lean_ctor_get(x_20, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_20); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_2, x_1); -if (x_11 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_12; +lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_10); -return x_12; +lean_dec(x_2); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_array_uget(x_3, x_2); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_3, x_2, x_14); -x_16 = lean_box(0); -x_17 = lean_ctor_get(x_13, 3); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_array_uget(x_5, x_4); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_5, x_4, x_16); x_18 = lean_box(0); +x_19 = lean_ctor_get(x_15, 3); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_16, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_19) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_18, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = 1; -x_23 = lean_usize_add(x_2, x_22); -x_24 = lean_array_uset(x_15, x_2, x_20); -x_2 = x_23; -x_3 = x_24; -x_10 = x_21; +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = 1; +x_25 = lean_usize_add(x_4, x_24); +x_26 = lean_array_uset(x_17, x_4, x_22); +x_4 = x_25; +x_5 = x_26; +x_12 = x_23; goto _start; } else { -uint8_t x_26; -lean_dec(x_15); +uint8_t x_28; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) { -return x_19; +return x_21; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_19, 0); -x_28 = lean_ctor_get(x_19, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_19); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_21); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_17, 0); -lean_inc(x_30); -lean_dec(x_17); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ lean_object* x_32; lean_object* x_33; -lean_dec(x_30); -x_32 = lean_box(0); +x_32 = lean_ctor_get(x_19, 0); +lean_inc(x_32); +lean_dec(x_19); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_32); +x_34 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_16, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_33) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_18, x_34, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = 1; -x_37 = lean_usize_add(x_2, x_36); -x_38 = lean_array_uset(x_15, x_2, x_34); -x_2 = x_37; -x_3 = x_38; -x_10 = x_35; +lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; lean_object* x_40; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = 1; +x_39 = lean_usize_add(x_4, x_38); +x_40 = lean_array_uset(x_17, x_4, x_36); +x_4 = x_39; +x_5 = x_40; +x_12 = x_37; goto _start; } else { -uint8_t x_40; -lean_dec(x_15); +uint8_t x_42; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_40 = !lean_is_exclusive(x_33); -if (x_40 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_35); +if (x_42 == 0) { -return x_33; +return x_35; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_33, 0); -x_42 = lean_ctor_get(x_33, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_33); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_35, 0); +x_44 = lean_ctor_get(x_35, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_35); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_31); -if (x_44 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_33); +if (x_46 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_31, 0); -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -lean_dec(x_45); -x_47 = l_Lean_Language_SnapshotTask_get___rarg(x_46); -if (lean_obj_tag(x_47) == 0) +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_33, 0); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Language_SnapshotTask_get___rarg(x_48); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_48; lean_object* x_49; -lean_free_object(x_31); -lean_dec(x_30); -x_48 = lean_box(0); +lean_object* x_50; lean_object* x_51; +lean_free_object(x_33); +lean_dec(x_32); +x_50 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_49 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_16, x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_49) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_51 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_18, x_50, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = 1; -x_53 = lean_usize_add(x_2, x_52); -x_54 = lean_array_uset(x_15, x_2, x_50); -x_2 = x_53; -x_3 = x_54; -x_10 = x_51; +lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = 1; +x_55 = lean_usize_add(x_4, x_54); +x_56 = lean_array_uset(x_17, x_4, x_52); +x_4 = x_55; +x_5 = x_56; +x_12 = x_53; goto _start; } else { -uint8_t x_56; -lean_dec(x_15); +uint8_t x_58; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_56 = !lean_is_exclusive(x_49); -if (x_56 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_51); +if (x_58 == 0) { -return x_49; +return x_51; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_49, 0); -x_58 = lean_ctor_get(x_49, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_49); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_51, 0); +x_60 = lean_ctor_get(x_51, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_51); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_47); -if (x_60 == 0) +uint8_t x_62; +x_62 = !lean_is_exclusive(x_49); +if (x_62 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_47, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_30, 1); -lean_inc(x_62); -lean_dec(x_30); -x_63 = lean_io_promise_resolve(x_47, x_62, x_10); -lean_dec(x_62); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = lean_ctor_get(x_49, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_32, 1); +lean_inc(x_64); +lean_dec(x_32); +x_65 = lean_io_promise_resolve(x_49, x_64, x_12); +lean_dec(x_64); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_65 = lean_ctor_get(x_63, 1); -x_66 = lean_ctor_get(x_63, 0); -lean_dec(x_66); -x_67 = lean_ctor_get(x_61, 1); -lean_inc(x_67); -x_68 = lean_ctor_get(x_61, 2); -lean_inc(x_68); -lean_dec(x_61); -lean_ctor_set(x_63, 1, x_67); -lean_ctor_set(x_63, 0, x_68); -lean_ctor_set(x_31, 0, x_63); -x_69 = lean_box(0); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_67 = lean_ctor_get(x_65, 1); +x_68 = lean_ctor_get(x_65, 0); +lean_dec(x_68); +x_69 = lean_ctor_get(x_63, 1); +lean_inc(x_69); +x_70 = lean_ctor_get(x_63, 2); +lean_inc(x_70); +lean_dec(x_63); +lean_ctor_set(x_65, 1, x_69); +lean_ctor_set(x_65, 0, x_70); +lean_ctor_set(x_33, 0, x_65); +x_71 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_31, x_69, x_4, x_5, x_6, x_7, x_8, x_9, x_65); -if (lean_obj_tag(x_70) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_72 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_33, x_71, x_6, x_7, x_8, x_9, x_10, x_11, x_67); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); -x_73 = 1; -x_74 = lean_usize_add(x_2, x_73); -x_75 = lean_array_uset(x_15, x_2, x_71); -x_2 = x_74; -x_3 = x_75; -x_10 = x_72; +lean_object* x_73; lean_object* x_74; size_t x_75; size_t x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = 1; +x_76 = lean_usize_add(x_4, x_75); +x_77 = lean_array_uset(x_17, x_4, x_73); +x_4 = x_76; +x_5 = x_77; +x_12 = x_74; goto _start; } else { -uint8_t x_77; -lean_dec(x_15); +uint8_t x_79; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_77 = !lean_is_exclusive(x_70); -if (x_77 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_79 = !lean_is_exclusive(x_72); +if (x_79 == 0) { -return x_70; +return x_72; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_70, 0); -x_79 = lean_ctor_get(x_70, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_70); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_72, 0); +x_81 = lean_ctor_get(x_72, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_72); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_81 = lean_ctor_get(x_63, 1); -lean_inc(x_81); -lean_dec(x_63); -x_82 = lean_ctor_get(x_61, 1); -lean_inc(x_82); -x_83 = lean_ctor_get(x_61, 2); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_83 = lean_ctor_get(x_65, 1); lean_inc(x_83); -lean_dec(x_61); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_82); -lean_ctor_set(x_31, 0, x_84); -x_85 = lean_box(0); +lean_dec(x_65); +x_84 = lean_ctor_get(x_63, 1); +lean_inc(x_84); +x_85 = lean_ctor_get(x_63, 2); +lean_inc(x_85); +lean_dec(x_63); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_84); +lean_ctor_set(x_33, 0, x_86); +x_87 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_86 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_31, x_85, x_4, x_5, x_6, x_7, x_8, x_9, x_81); -if (lean_obj_tag(x_86) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_33, x_87, x_6, x_7, x_8, x_9, x_10, x_11, x_83); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_87; lean_object* x_88; size_t x_89; size_t x_90; lean_object* x_91; -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = 1; -x_90 = lean_usize_add(x_2, x_89); -x_91 = lean_array_uset(x_15, x_2, x_87); -x_2 = x_90; -x_3 = x_91; -x_10 = x_88; +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_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = 1; +x_92 = lean_usize_add(x_4, x_91); +x_93 = lean_array_uset(x_17, x_4, x_89); +x_4 = x_92; +x_5 = x_93; +x_12 = x_90; goto _start; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_15); +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_93 = lean_ctor_get(x_86, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_86, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_95 = x_86; +lean_dec(x_2); +lean_dec(x_1); +x_95 = lean_ctor_get(x_88, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_88, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_97 = x_88; } else { - lean_dec_ref(x_86); - x_95 = lean_box(0); + lean_dec_ref(x_88); + x_97 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(1, 2, 0); } else { - x_96 = x_95; + x_98 = x_97; } -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; } } } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_97 = lean_ctor_get(x_47, 0); -lean_inc(x_97); -lean_dec(x_47); -lean_inc(x_97); -x_98 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_98, 0, x_97); -x_99 = lean_ctor_get(x_30, 1); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_99 = lean_ctor_get(x_49, 0); lean_inc(x_99); -lean_dec(x_30); -x_100 = lean_io_promise_resolve(x_98, x_99, x_10); -lean_dec(x_99); -x_101 = lean_ctor_get(x_100, 1); +lean_dec(x_49); +lean_inc(x_99); +x_100 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_100, 0, x_99); +x_101 = lean_ctor_get(x_32, 1); lean_inc(x_101); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - lean_ctor_release(x_100, 1); - x_102 = x_100; +lean_dec(x_32); +x_102 = lean_io_promise_resolve(x_100, x_101, x_12); +lean_dec(x_101); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_104 = x_102; } else { - lean_dec_ref(x_100); - x_102 = lean_box(0); + lean_dec_ref(x_102); + x_104 = lean_box(0); } -x_103 = lean_ctor_get(x_97, 1); -lean_inc(x_103); -x_104 = lean_ctor_get(x_97, 2); -lean_inc(x_104); -lean_dec(x_97); -if (lean_is_scalar(x_102)) { - x_105 = lean_alloc_ctor(0, 2, 0); +x_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); +x_106 = lean_ctor_get(x_99, 2); +lean_inc(x_106); +lean_dec(x_99); +if (lean_is_scalar(x_104)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_105 = x_102; + x_107 = x_104; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_103); -lean_ctor_set(x_31, 0, x_105); -x_106 = lean_box(0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +lean_ctor_set(x_33, 0, x_107); +x_108 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_107 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_31, x_106, x_4, x_5, x_6, x_7, x_8, x_9, x_101); -if (lean_obj_tag(x_107) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_33, x_108, x_6, x_7, x_8, x_9, x_10, x_11, x_103); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_108; lean_object* x_109; size_t x_110; size_t x_111; lean_object* x_112; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = 1; -x_111 = lean_usize_add(x_2, x_110); -x_112 = lean_array_uset(x_15, x_2, x_108); -x_2 = x_111; -x_3 = x_112; -x_10 = x_109; +lean_object* x_110; lean_object* x_111; size_t x_112; size_t x_113; lean_object* x_114; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = 1; +x_113 = lean_usize_add(x_4, x_112); +x_114 = lean_array_uset(x_17, x_4, x_110); +x_4 = x_113; +x_5 = x_114; +x_12 = x_111; goto _start; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_15); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_114 = lean_ctor_get(x_107, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_107, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_116 = x_107; +lean_dec(x_2); +lean_dec(x_1); +x_116 = lean_ctor_get(x_109, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_109, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_118 = x_109; } else { - lean_dec_ref(x_107); - x_116 = lean_box(0); + lean_dec_ref(x_109); + x_118 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_119 = x_118; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; } } } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_31, 0); -lean_inc(x_118); -lean_dec(x_31); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -lean_dec(x_118); -x_120 = l_Lean_Language_SnapshotTask_get___rarg(x_119); -if (lean_obj_tag(x_120) == 0) +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_33, 0); +lean_inc(x_120); +lean_dec(x_33); +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +x_122 = l_Lean_Language_SnapshotTask_get___rarg(x_121); +if (lean_obj_tag(x_122) == 0) { -lean_object* x_121; lean_object* x_122; -lean_dec(x_30); -x_121 = lean_box(0); +lean_object* x_123; lean_object* x_124; +lean_dec(x_32); +x_123 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_122 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_16, x_121, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_122) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_124 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_18, x_123, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_124) == 0) { -lean_object* x_123; lean_object* x_124; size_t x_125; size_t x_126; lean_object* x_127; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = 1; -x_126 = lean_usize_add(x_2, x_125); -x_127 = lean_array_uset(x_15, x_2, x_123); -x_2 = x_126; -x_3 = x_127; -x_10 = x_124; +lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +x_127 = 1; +x_128 = lean_usize_add(x_4, x_127); +x_129 = lean_array_uset(x_17, x_4, x_125); +x_4 = x_128; +x_5 = x_129; +x_12 = x_126; goto _start; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_15); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_129 = lean_ctor_get(x_122, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_122, 1); -lean_inc(x_130); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_131 = x_122; +lean_dec(x_2); +lean_dec(x_1); +x_131 = lean_ctor_get(x_124, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_124, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_133 = x_124; } else { - lean_dec_ref(x_122); - x_131 = lean_box(0); + lean_dec_ref(x_124); + x_133 = lean_box(0); } -if (lean_is_scalar(x_131)) { - x_132 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(1, 2, 0); } else { - x_132 = x_131; + x_134 = x_133; } -lean_ctor_set(x_132, 0, x_129); -lean_ctor_set(x_132, 1, x_130); -return x_132; +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; } } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_133 = lean_ctor_get(x_120, 0); -lean_inc(x_133); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - x_134 = x_120; +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_135 = lean_ctor_get(x_122, 0); +lean_inc(x_135); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + x_136 = x_122; } else { - lean_dec_ref(x_120); - x_134 = lean_box(0); + lean_dec_ref(x_122); + x_136 = lean_box(0); } -lean_inc(x_133); -if (lean_is_scalar(x_134)) { - x_135 = lean_alloc_ctor(1, 1, 0); +lean_inc(x_135); +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 1, 0); } else { - x_135 = x_134; + x_137 = x_136; } -lean_ctor_set(x_135, 0, x_133); -x_136 = lean_ctor_get(x_30, 1); -lean_inc(x_136); -lean_dec(x_30); -x_137 = lean_io_promise_resolve(x_135, x_136, x_10); -lean_dec(x_136); -x_138 = lean_ctor_get(x_137, 1); +lean_ctor_set(x_137, 0, x_135); +x_138 = lean_ctor_get(x_32, 1); lean_inc(x_138); -if (lean_is_exclusive(x_137)) { - lean_ctor_release(x_137, 0); - lean_ctor_release(x_137, 1); - x_139 = x_137; +lean_dec(x_32); +x_139 = lean_io_promise_resolve(x_137, x_138, x_12); +lean_dec(x_138); +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_141 = x_139; } else { - lean_dec_ref(x_137); - x_139 = lean_box(0); + lean_dec_ref(x_139); + x_141 = lean_box(0); } -x_140 = lean_ctor_get(x_133, 1); -lean_inc(x_140); -x_141 = lean_ctor_get(x_133, 2); -lean_inc(x_141); -lean_dec(x_133); -if (lean_is_scalar(x_139)) { - x_142 = lean_alloc_ctor(0, 2, 0); +x_142 = lean_ctor_get(x_135, 1); +lean_inc(x_142); +x_143 = lean_ctor_get(x_135, 2); +lean_inc(x_143); +lean_dec(x_135); +if (lean_is_scalar(x_141)) { + x_144 = lean_alloc_ctor(0, 2, 0); } else { - x_142 = x_139; + x_144 = x_141; } -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_140); -x_143 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_143, 0, x_142); -x_144 = lean_box(0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_142); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_144); +x_146 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_145 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_13, x_16, x_143, x_144, x_4, x_5, x_6, x_7, x_8, x_9, x_138); -if (lean_obj_tag(x_145) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_147 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_15, x_1, x_2, x_18, x_145, x_146, x_6, x_7, x_8, x_9, x_10, x_11, x_140); +if (lean_obj_tag(x_147) == 0) { -lean_object* x_146; lean_object* x_147; size_t x_148; size_t x_149; lean_object* x_150; -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = 1; -x_149 = lean_usize_add(x_2, x_148); -x_150 = lean_array_uset(x_15, x_2, x_146); -x_2 = x_149; -x_3 = x_150; -x_10 = x_147; +lean_object* x_148; lean_object* x_149; size_t x_150; size_t x_151; lean_object* x_152; +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = 1; +x_151 = lean_usize_add(x_4, x_150); +x_152 = lean_array_uset(x_17, x_4, x_148); +x_4 = x_151; +x_5 = x_152; +x_12 = x_149; goto _start; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -lean_dec(x_15); +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_152 = lean_ctor_get(x_145, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_145, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_154 = x_145; +lean_dec(x_2); +lean_dec(x_1); +x_154 = lean_ctor_get(x_147, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_147, 1); +lean_inc(x_155); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_156 = x_147; } else { - lean_dec_ref(x_145); - x_154 = lean_box(0); + lean_dec_ref(x_147); + x_156 = lean_box(0); } -if (lean_is_scalar(x_154)) { - x_155 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_156)) { + x_157 = lean_alloc_ctor(1, 2, 0); } else { - x_155 = x_154; + x_157 = x_156; } -lean_ctor_set(x_155, 0, x_152); -lean_ctor_set(x_155, 1, x_153); -return x_155; +lean_ctor_set(x_157, 0, x_154); +lean_ctor_set(x_157, 1, x_155); +return x_157; } } } @@ -22134,14 +25312,14 @@ return x_155; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_array_size(x_1); -x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(x_9, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_array_size(x_1); +x_12 = 0; +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15(x_2, x_3, x_11, x_12, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { @@ -22158,57 +25336,105 @@ lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_14; -x_14 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_11; -x_11 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_7); lean_dec(x_7); -return x_11; +x_18 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_15; -x_15 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); -return x_15; +lean_object* x_17; +x_17 = l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_15; -x_15 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__7___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); lean_dec(x_9); -return x_15; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_15; -x_15 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_7); lean_dec(x_7); -return x_15; +x_18 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_19; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -22219,25 +25445,360 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -lean_object* x_12; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_4); -return x_12; +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_18; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_12 = lean_unbox_usize(x_2); +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_8); +lean_dec(x_7); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_7); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_11); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_6); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; } } LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { @@ -22312,45 +25873,6 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_2, x_3); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; -lean_dec(x_4); -x_12 = lean_array_uget(x_1, x_2); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 5); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Expr_collectFVars(x_14, x_5, x_6, x_7, x_8, x_9, x_10); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = 1; -x_19 = lean_usize_add(x_2, x_18); -x_2 = x_19; -x_4 = x_16; -x_10 = x_17; -goto _start; -} -else -{ -lean_object* x_21; -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_4); -lean_ctor_set(x_21, 1, x_10); -return x_21; -} -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -22381,7 +25903,7 @@ x_27 = 0; x_28 = lean_usize_of_nat(x_10); lean_dec(x_10); x_29 = lean_box(0); -x_30 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3(x_1, x_27, x_28, x_29, x_4, x_5, x_6, x_7, x_8, x_9); +x_30 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5(x_1, x_27, x_28, x_29, x_4, x_5, x_6, x_7, x_8, x_9); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); @@ -22461,24 +25983,6 @@ lean_dec(x_1); return x_13; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_12 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_13 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -return x_13; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -22494,34 +25998,11 @@ lean_dec(x_1); return x_10; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(8u); -x_2 = l_Lean_mkHashSetImp___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; -x_3 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_1); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_13 = lean_st_mk_ref(x_12, x_11); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); @@ -24424,7 +27905,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_2 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24461,7 +27942,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_array_uget(x_8, x_10); x_17 = lean_ctor_get(x_1, 1); lean_inc(x_17); -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_19 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; lean_inc(x_5); x_20 = lean_alloc_ctor(0, 3, 0); @@ -24589,7 +28070,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_array_uget(x_7, x_9); x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); -x_17 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_17 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_18 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; lean_inc(x_5); x_19 = lean_alloc_ctor(0, 3, 0); @@ -26662,7 +30143,7 @@ if (x_30 == 0) lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_31 = lean_ctor_get(x_29, 0); lean_dec(x_31); -x_32 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_32 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_33 = l_Lean_CollectFVars_main(x_16, x_32); x_34 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_1, x_33); lean_ctor_set(x_29, 0, x_34); @@ -26674,7 +30155,7 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean x_35 = lean_ctor_get(x_29, 1); lean_inc(x_35); lean_dec(x_29); -x_36 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_36 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_37 = l_Lean_CollectFVars_main(x_16, x_36); x_38 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_1, x_37); x_39 = lean_alloc_ctor(0, 2, 0); @@ -26724,7 +30205,7 @@ if (lean_is_exclusive(x_50)) { lean_dec_ref(x_50); x_52 = lean_box(0); } -x_53 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_53 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_54 = l_Lean_CollectFVars_main(x_16, x_53); x_55 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_1, x_54); if (lean_is_scalar(x_52)) { @@ -27342,7 +30823,7 @@ x_70 = lean_st_ref_set(x_4, x_56, x_57); x_71 = lean_ctor_get(x_70, 1); lean_inc(x_71); lean_dec(x_70); -x_72 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_72 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_73 = l_Lean_CollectFVars_main(x_50, x_72); x_74 = l_Lean_CollectFVars_main(x_53, x_73); x_75 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_12, x_74); @@ -27387,7 +30868,7 @@ x_90 = lean_st_ref_set(x_4, x_89, x_57); x_91 = lean_ctor_get(x_90, 1); lean_inc(x_91); lean_dec(x_90); -x_92 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_92 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_93 = l_Lean_CollectFVars_main(x_50, x_92); x_94 = l_Lean_CollectFVars_main(x_53, x_93); x_95 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_12, x_94); @@ -27616,7 +31097,7 @@ x_146 = lean_st_ref_set(x_4, x_145, x_130); x_147 = lean_ctor_get(x_146, 1); lean_inc(x_147); lean_dec(x_146); -x_148 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; +x_148 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2; x_149 = l_Lean_CollectFVars_main(x_123, x_148); x_150 = l_Lean_CollectFVars_main(x_126, x_149); x_151 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(x_12, x_150); @@ -29423,7 +32904,7 @@ else { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; x_18 = lean_array_uget(x_5, x_7); -x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_20 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; lean_inc(x_3); x_21 = lean_alloc_ctor(0, 3, 0); @@ -29484,7 +32965,7 @@ 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_34; size_t x_35; size_t x_36; x_17 = lean_array_uget(x_4, x_6); -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_19 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; lean_inc(x_2); x_20 = lean_alloc_ctor(0, 3, 0); @@ -39428,7 +42909,7 @@ static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels___closed__1( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; +x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1; x_2 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -41014,7 +44495,45 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; +x_5 = lean_array_uget(x_1, x_2); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get_uint8(x_6, sizeof(void*)*9); +lean_dec(x_6); +x_8 = l_Lean_Elab_DefKind_isTheorem(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = 1; +return x_9; +} +else +{ +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_2 = x_11; +goto _start; +} +} +else +{ +uint8_t x_13; +x_13 = 0; +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -41095,7 +44614,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; @@ -41330,7 +44849,7 @@ return x_62; } } } -static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1() { +static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1() { _start: { lean_object* x_1; @@ -41338,7 +44857,7 @@ x_1 = l_Lean_declRangeExt; return x_1; } } -static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2() { +static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -41351,12 +44870,12 @@ lean_ctor_set(x_2, 3, x_1); return x_2; } } -static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3() { +static lean_object* _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__13; -x_2 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2; +x_2 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2; x_3 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_1); @@ -41368,7 +44887,7 @@ lean_ctor_set(x_3, 6, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -41385,7 +44904,7 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_11, 0); x_15 = lean_ctor_get(x_11, 4); lean_dec(x_15); -x_16 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1; +x_16 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1; x_17 = l_Lean_MapDeclarationExtension_insert___rarg(x_16, x_14, x_1, x_2); x_18 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__14; lean_ctor_set(x_11, 4, x_18); @@ -41406,7 +44925,7 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; x_25 = lean_ctor_get(x_22, 1); lean_dec(x_25); -x_26 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3; +x_26 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3; lean_ctor_set(x_22, 1, x_26); x_27 = lean_st_ref_set(x_6, x_22, x_23); x_28 = !lean_is_exclusive(x_27); @@ -41444,7 +44963,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_22); -x_38 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3; +x_38 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3; x_39 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_39, 0, x_34); lean_ctor_set(x_39, 1, x_38); @@ -41489,7 +45008,7 @@ lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); lean_dec(x_11); -x_51 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1; +x_51 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1; x_52 = l_Lean_MapDeclarationExtension_insert___rarg(x_51, x_45, x_1, x_2); x_53 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__14; x_54 = lean_alloc_ctor(0, 7, 0); @@ -41529,7 +45048,7 @@ if (lean_is_exclusive(x_58)) { lean_dec_ref(x_58); x_64 = lean_box(0); } -x_65 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3; +x_65 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3; if (lean_is_scalar(x_64)) { x_66 = lean_alloc_ctor(0, 5, 0); } else { @@ -41563,7 +45082,7 @@ return x_71; } } } -static lean_object* _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1() { +static lean_object* _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1() { _start: { lean_object* x_1; @@ -41571,32 +45090,32 @@ x_1 = lean_mk_string_unchecked("example", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2() { +static lean_object* _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__1; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__2; x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__3; -x_4 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1; +x_4 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_2); x_10 = l_Lean_Syntax_getKind(x_2); -x_11 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2; +x_11 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2; x_12 = lean_name_eq(x_10, x_11); lean_dec(x_10); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_inc(x_7); -x_13 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -41604,7 +45123,7 @@ lean_inc(x_15); lean_dec(x_13); x_16 = l_Lean_Elab_getDeclarationSelectionRef(x_2); lean_inc(x_7); -x_17 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +x_17 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_15); lean_dec(x_16); x_18 = !lean_is_exclusive(x_17); if (x_18 == 0) @@ -41614,7 +45133,7 @@ x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_17, 1); lean_ctor_set(x_17, 1, x_19); lean_ctor_set(x_17, 0, x_14); -x_21 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_21 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_20); lean_dec(x_7); return x_21; } @@ -41629,7 +45148,7 @@ lean_dec(x_17); x_24 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_24, 0, x_14); lean_ctor_set(x_24, 1, x_22); -x_25 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_1, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_23); +x_25 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9(x_1, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_23); lean_dec(x_7); return x_25; } @@ -41648,7 +45167,7 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; @@ -41714,7 +45233,7 @@ x_29 = lean_ctor_get(x_14, 0); lean_inc(x_29); lean_dec(x_14); lean_inc(x_9); -x_30 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6(x_28, x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_30 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_28, x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); @@ -41746,7 +45265,7 @@ x_41 = lean_ctor_get(x_14, 0); lean_inc(x_41); lean_dec(x_14); lean_inc(x_9); -x_42 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6(x_40, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_42 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_40, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11); x_43 = lean_ctor_get(x_42, 1); lean_inc(x_43); lean_dec(x_42); @@ -41761,7 +45280,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -41795,7 +45314,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -41830,7 +45349,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -41865,7 +45384,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; @@ -41876,7 +45395,7 @@ lean_dec(x_12); x_14 = l_List_toArrayAux___rarg(x_11, x_13); x_15 = lean_array_size(x_14); x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11(x_15, x_16, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12(x_15, x_16, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -41900,7 +45419,7 @@ x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); x_24 = lean_box(0); -x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12(x_1, x_18, x_20, x_16, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(x_1, x_18, x_20, x_16, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_23); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -41938,7 +45457,7 @@ x_31 = lean_ctor_get(x_21, 1); lean_inc(x_31); lean_dec(x_21); x_32 = lean_box(0); -x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(x_1, x_18, x_20, x_16, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14(x_1, x_18, x_20, x_16, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -41970,7 +45489,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -42004,7 +45523,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -42039,7 +45558,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -42074,7 +45593,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; @@ -42085,7 +45604,7 @@ lean_dec(x_12); x_14 = l_List_toArrayAux___rarg(x_11, x_13); x_15 = lean_array_size(x_14); x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15(x_15, x_16, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16(x_15, x_16, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -42109,7 +45628,7 @@ x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); x_24 = lean_box(0); -x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16(x_1, x_18, x_20, x_16, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(x_1, x_18, x_20, x_16, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_23); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -42147,7 +45666,7 @@ x_31 = lean_ctor_get(x_21, 1); lean_inc(x_31); lean_dec(x_21); x_32 = lean_box(0); -x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(x_1, x_18, x_20, x_16, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18(x_1, x_18, x_20, x_16, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -42486,236 +46005,309 @@ return x_68; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -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_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_16 = lean_array_size(x_1); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3(x_16, x_2, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +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_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_17 = lean_array_size(x_1); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3(x_17, x_2, x_1, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Elab_Term_getLetRecsToLift___rarg(x_10, x_11, x_12, x_13, x_14, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Term_getLetRecsToLift___rarg(x_11, x_12, x_13, x_14, x_15, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_box(0); -x_24 = l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__2(x_21, x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_22); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_box(0); +x_25 = l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__2(x_22, x_24, x_10, x_11, x_12, x_13, x_14, x_15, x_23); +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_13); -lean_inc(x_11); -lean_inc(x_9); -lean_inc_n(x_25, 2); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_10); +lean_inc_n(x_26, 2); lean_inc(x_3); -x_27 = l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1(x_3, x_25, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_26); -if (lean_obj_tag(x_27) == 0) +x_28 = l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1(x_3, x_26, x_26, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_box_usize(x_2); -lean_inc(x_25); -lean_inc(x_8); -lean_inc(x_18); -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__1___boxed), 16, 8); -lean_closure_set(x_30, 0, x_18); -lean_closure_set(x_30, 1, x_3); -lean_closure_set(x_30, 2, x_8); -lean_closure_set(x_30, 3, x_25); -lean_closure_set(x_30, 4, x_29); -lean_closure_set(x_30, 5, x_4); -lean_closure_set(x_30, 6, x_5); -lean_closure_set(x_30, 7, x_6); -x_31 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___rarg(x_7, x_18, x_8, x_25, x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_28); -lean_dec(x_8); -lean_dec(x_18); -return x_31; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_ctor_get(x_14, 2); +lean_inc(x_30); +x_31 = lean_array_get_size(x_19); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_lt(x_32, x_31); +x_34 = lean_box_usize(x_2); +lean_inc(x_26); +lean_inc(x_9); +lean_inc(x_19); +x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__1___boxed), 16, 8); +lean_closure_set(x_35, 0, x_19); +lean_closure_set(x_35, 1, x_3); +lean_closure_set(x_35, 2, x_9); +lean_closure_set(x_35, 3, x_26); +lean_closure_set(x_35, 4, x_34); +lean_closure_set(x_35, 5, x_4); +lean_closure_set(x_35, 6, x_5); +lean_closure_set(x_35, 7, x_6); +if (x_33 == 0) +{ +lean_object* x_36; uint8_t x_37; +lean_dec(x_31); +x_36 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1; +x_37 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_30, x_36); +lean_dec(x_30); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_26); +lean_dec(x_9); +x_38 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(x_7, x_8, x_19, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_19); +return x_38; } else { -uint8_t x_32; -lean_dec(x_25); -lean_dec(x_18); +lean_object* x_39; +x_39 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___rarg(x_7, x_19, x_9, x_26, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_9); +lean_dec(x_19); +return x_39; +} +} +else +{ +size_t x_40; uint8_t x_41; +x_40 = lean_usize_of_nat(x_31); +lean_dec(x_31); +x_41 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5(x_19, x_2, x_40); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +x_42 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1; +x_43 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_30, x_42); +lean_dec(x_30); +if (x_43 == 0) +{ +lean_object* x_44; +lean_dec(x_26); +lean_dec(x_9); +x_44 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___rarg(x_7, x_8, x_19, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_19); +return x_44; +} +else +{ +lean_object* x_45; +x_45 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___rarg(x_7, x_19, x_9, x_26, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_9); +lean_dec(x_19); +return x_45; +} +} +else +{ +lean_object* x_46; +lean_dec(x_30); +x_46 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___rarg(x_7, x_19, x_9, x_26, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +lean_dec(x_9); +lean_dec(x_19); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_19); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +x_47 = !lean_is_exclusive(x_28); +if (x_47 == 0) { -return x_27; +return x_28; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_28, 0); +x_49 = lean_ctor_get(x_28, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_28); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; -x_14 = lean_array_get_size(x_6); -x_15 = lean_unsigned_to_nat(0u); -lean_inc(x_6); -x_16 = l_Array_toSubarray___rarg(x_6, x_15, x_14); -x_17 = lean_array_size(x_1); -x_18 = 0; +lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_15 = lean_array_get_size(x_7); +x_16 = lean_unsigned_to_nat(0u); +lean_inc(x_7); +x_17 = l_Array_toSubarray___rarg(x_7, x_16, x_15); +x_18 = lean_array_size(x_1); +x_19 = 0; +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_19 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__1(x_1, x_17, x_18, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) +x_20 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__1(x_1, x_18, x_19, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_44; -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - x_21 = x_19; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_45; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_22 = x_20; } else { - lean_dec_ref(x_19); - x_21 = lean_box(0); + lean_dec_ref(x_20); + x_22 = lean_box(0); } +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); lean_inc(x_2); -x_44 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_20); -if (lean_obj_tag(x_44) == 0) +x_45 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(x_2, x_5, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_45; lean_object* x_46; uint8_t x_47; uint8_t x_48; lean_object* x_49; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); +lean_object* x_46; lean_object* x_47; uint8_t x_48; uint8_t x_49; lean_object* x_50; +x_46 = lean_ctor_get(x_45, 0); lean_inc(x_46); -lean_dec(x_44); -x_47 = 1; -x_48 = 0; +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = 1; +x_49 = 0; +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_49 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_47, x_48, x_7, x_8, x_9, x_10, x_11, x_12, x_46); -if (lean_obj_tag(x_49) == 0) +lean_inc(x_8); +x_50 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_48, x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_47); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_50; size_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_21); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -lean_dec(x_49); -x_51 = lean_array_size(x_45); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(x_51, x_18, x_45, x_7, x_8, x_9, x_10, x_11, x_12, x_50); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); +lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_22); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = lean_array_size(x_46); +x_53 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(x_52, x_19, x_46, x_8, x_9, x_10, x_11, x_12, x_13, x_51); +x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); -lean_dec(x_52); -x_55 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_18, x_6, x_3, x_4, x_1, x_5, x_53, x_7, x_8, x_9, x_10, x_11, x_12, x_54); -return x_55; +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_19, x_7, x_3, x_4, x_1, x_5, x_6, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_55); +lean_dec(x_6); +lean_dec(x_5); +return x_56; } else { -lean_object* x_56; lean_object* x_57; -lean_dec(x_45); -x_56 = lean_ctor_get(x_49, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_49, 1); +lean_object* x_57; lean_object* x_58; +lean_dec(x_46); +x_57 = lean_ctor_get(x_50, 0); lean_inc(x_57); -lean_dec(x_49); -x_22 = x_56; +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_dec(x_50); x_23 = x_57; -goto block_43; +x_24 = x_58; +goto block_44; } } else { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_44, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_44, 1); +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_45, 0); lean_inc(x_59); -lean_dec(x_44); -x_22 = x_58; +x_60 = lean_ctor_get(x_45, 1); +lean_inc(x_60); +lean_dec(x_45); x_23 = x_59; -goto block_43; +x_24 = x_60; +goto block_44; } -block_43: -{ -uint8_t x_24; -x_24 = l_Lean_Exception_isInterrupt(x_22); -if (x_24 == 0) +block_44: { uint8_t x_25; -x_25 = l_Lean_Exception_isRuntime(x_22); +x_25 = l_Lean_Exception_isInterrupt(x_23); if (x_25 == 0) { -lean_object* x_26; -lean_dec(x_21); -lean_inc(x_11); -x_26 = l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_23); -if (lean_obj_tag(x_26) == 0) +uint8_t x_26; +x_26 = l_Lean_Exception_isRuntime(x_23); +if (x_26 == 0) { -lean_object* x_27; size_t x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_array_size(x_2); +lean_object* x_27; +lean_dec(x_22); +lean_inc(x_12); +x_27 = l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(x_23, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; size_t x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_array_size(x_2); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); lean_inc(x_2); -x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5(x_28, x_18, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_27); -if (lean_obj_tag(x_29) == 0) +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6(x_29, x_19, x_2, x_8, x_9, x_10, x_11, x_12, x_13, x_28); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_18, x_6, x_3, x_4, x_1, x_5, x_30, x_7, x_8, x_9, x_10, x_11, x_12, x_31); -return x_32; +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_19, x_7, x_3, x_4, x_1, x_5, x_6, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_32); +lean_dec(x_6); +lean_dec(x_5); +return x_33; } else { -uint8_t x_33; +uint8_t x_34; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -42723,33 +46315,35 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_29); -if (x_33 == 0) +x_34 = !lean_is_exclusive(x_30); +if (x_34 == 0) { -return x_29; +return x_30; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_29, 0); -x_35 = lean_ctor_get(x_29, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_30, 0); +x_36 = lean_ctor_get(x_30, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_29); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_30); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } else { -uint8_t x_37; +uint8_t x_38; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -42757,33 +46351,35 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_26); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_27); +if (x_38 == 0) { -return x_26; +return x_27; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_26, 0); -x_39 = lean_ctor_get(x_26, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_27, 0); +x_40 = lean_ctor_get(x_27, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_26); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_27); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } else { -lean_object* x_41; +lean_object* x_42; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -42791,24 +46387,26 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_21)) { - x_41 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_22)) { + x_42 = lean_alloc_ctor(1, 2, 0); } else { - x_41 = x_21; - lean_ctor_set_tag(x_41, 1); + x_42 = x_22; + lean_ctor_set_tag(x_42, 1); } -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_23); -return x_41; +lean_ctor_set(x_42, 0, x_23); +lean_ctor_set(x_42, 1, x_24); +return x_42; } } else { -lean_object* x_42; +lean_object* x_43; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -42816,25 +46414,27 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (lean_is_scalar(x_21)) { - x_42 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_22)) { + x_43 = lean_alloc_ctor(1, 2, 0); } else { - x_42 = x_21; - lean_ctor_set_tag(x_42, 1); + x_43 = x_22; + lean_ctor_set_tag(x_43, 1); } -lean_ctor_set(x_42, 0, x_22); -lean_ctor_set(x_42, 1, x_23); -return x_42; +lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 1, x_24); +return x_43; } } } else { -uint8_t x_60; +uint8_t x_61; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -42842,252 +46442,258 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_19); -if (x_60 == 0) +x_61 = !lean_is_exclusive(x_20); +if (x_61 == 0) { -return x_19; +return x_20; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_19, 0); -x_62 = lean_ctor_get(x_19, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_20, 0); +x_63 = lean_ctor_get(x_20, 1); +lean_inc(x_63); lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_19); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; +lean_dec(x_20); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l_Lean_Elab_Term_getLevelNames___rarg(x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Elab_Term_getLevelNames___rarg(x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); lean_inc(x_1); -x_15 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders(x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders(x_1, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -if (lean_obj_tag(x_18) == 0) +x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders(x_1, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getAllUserLevelNames(x_20); lean_inc(x_20); -lean_dec(x_18); -x_21 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getAllUserLevelNames(x_19); -lean_inc(x_19); lean_inc(x_1); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__3___boxed), 13, 5); -lean_closure_set(x_22, 0, x_1); -lean_closure_set(x_22, 1, x_19); -lean_closure_set(x_22, 2, x_21); -lean_closure_set(x_22, 3, x_13); -lean_closure_set(x_22, 4, x_3); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__3), 14, 6); +lean_closure_set(x_23, 0, x_1); +lean_closure_set(x_23, 1, x_20); +lean_closure_set(x_23, 2, x_22); +lean_closure_set(x_23, 3, x_14); +lean_closure_set(x_23, 4, x_3); +lean_closure_set(x_23, 5, x_4); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_19); -x_23 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls___rarg(x_19, x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_20); -if (lean_obj_tag(x_23) == 0) +lean_inc(x_20); +x_24 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls___rarg(x_20, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_array_get_size(x_19); -x_26 = lean_unsigned_to_nat(0u); -x_27 = l_Array_toSubarray___rarg(x_19, x_26, x_25); -x_28 = lean_array_size(x_1); -x_29 = 0; -x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9(x_1, x_28, x_29, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -lean_dec(x_10); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; uint8_t x_32; +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_array_get_size(x_20); +x_27 = lean_unsigned_to_nat(0u); +x_28 = l_Array_toSubarray___rarg(x_20, x_27, x_26); +x_29 = lean_array_size(x_1); +x_30 = 0; +x_31 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10(x_1, x_29, x_30, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 0); -lean_dec(x_32); -x_33 = lean_box(0); -lean_ctor_set(x_30, 0, x_33); -return x_30; +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +x_34 = lean_box(0); +lean_ctor_set(x_31, 0, x_34); +return x_31; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_35); +lean_dec(x_31); +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +return x_37; } } else { -uint8_t x_37; -lean_dec(x_19); +uint8_t x_38; +lean_dec(x_20); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_23); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_24); +if (x_38 == 0) { -return x_23; +return x_24; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_23, 0); -x_39 = lean_ctor_get(x_23, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_24, 0); +x_40 = lean_ctor_get(x_24, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_23); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_24); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } else { -uint8_t x_41; -lean_dec(x_13); +uint8_t x_42; +lean_dec(x_14); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_41 = !lean_is_exclusive(x_18); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_19); +if (x_42 == 0) { -return x_18; +return x_19; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_18, 0); -x_43 = lean_ctor_get(x_18, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_19, 0); +x_44 = lean_ctor_get(x_19, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_18); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec(x_19); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_45; -lean_dec(x_13); +uint8_t x_46; +lean_dec(x_14); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_45 = !lean_is_exclusive(x_15); -if (x_45 == 0) +x_46 = !lean_is_exclusive(x_16); +if (x_46 == 0) { -return x_15; +return x_16; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_15, 0); -x_47 = lean_ctor_get(x_15, 1); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_16, 0); +x_48 = lean_ctor_get(x_16, 1); +lean_inc(x_48); lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_15); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_dec(x_16); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__4), 11, 3); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_4); -lean_closure_set(x_12, 2, x_2); -x_13 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; -x_14 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__10(x_13, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__4), 12, 4); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_5); +lean_closure_set(x_13, 2, x_2); +lean_closure_set(x_13, 3, x_3); +x_14 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; +x_15 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__11(x_14, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_box(0); -x_11 = lean_array_get_size(x_2); -lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__5), 11, 3); -lean_closure_set(x_12, 0, x_2); -lean_closure_set(x_12, 1, x_1); -lean_closure_set(x_12, 2, x_11); -x_13 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__14(x_10, x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_13; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_box(0); +x_12 = lean_array_get_size(x_3); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go___lambda__5), 12, 4); +lean_closure_set(x_13, 0, x_3); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_2); +lean_closure_set(x_13, 3, x_12); +x_14 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__15(x_11, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -43155,7 +46761,21 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMutualDef_go___spec__5(x_1, x_4, x_5); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -43163,17 +46783,17 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__5(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__6(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); @@ -43183,11 +46803,11 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -43197,11 +46817,11 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -43210,7 +46830,7 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -43218,7 +46838,7 @@ x_12 = lean_unbox_usize(x_2); lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__9(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__10(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -43228,7 +46848,7 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -43236,7 +46856,7 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__11(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__12(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -43246,7 +46866,7 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -43254,7 +46874,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__12(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -43265,7 +46885,7 @@ lean_dec(x_2); return x_15; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -43273,7 +46893,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__13(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -43284,7 +46904,7 @@ lean_dec(x_2); return x_15; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -43292,7 +46912,7 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__15(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef_go___spec__16(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -43302,7 +46922,7 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -43310,7 +46930,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__16(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -43321,7 +46941,7 @@ lean_dec(x_2); return x_15; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -43329,7 +46949,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__17(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__18(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -43350,24 +46970,16 @@ x_18 = l_Lean_Elab_Term_elabMutualDef_go___lambda__1(x_1, x_2, x_3, x_4, x_17, x return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -size_t x_16; lean_object* x_17; -x_16 = lean_unbox_usize(x_2); +size_t x_17; lean_object* x_18; +x_17 = lean_unbox_usize(x_2); lean_dec(x_2); -x_17 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_8); lean_dec(x_7); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_Elab_Term_elabMutualDef_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_5); -return x_14; +return x_18; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__5(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -45560,111 +49172,112 @@ x_10 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoCont return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_10; -x_10 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample(x_2); -if (x_10 == 0) +uint8_t x_11; +x_11 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample(x_3); +if (x_11 == 0) { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_elabMutualDef_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_12; +x_12 = l_Lean_Elab_Term_elabMutualDef_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go), 9, 2); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -x_13 = lean_st_ref_get(x_8, x_9); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabMutualDef_go), 10, 3); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_3); +x_14 = lean_st_ref_get(x_9, x_10); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 0); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___lambda__1___boxed), 7, 0); +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___lambda__1___boxed), 7, 0); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_18 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__2(x_12, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_15); -if (lean_obj_tag(x_18) == 0) +x_19 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__2(x_13, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_19); -return x_21; +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_20); +return x_22; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_19); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_18, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_19, 0); lean_inc(x_27); -lean_dec(x_18); -x_28 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_dec(x_19); +x_29 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -lean_ctor_set_tag(x_28, 1); -lean_ctor_set(x_28, 0, x_26); -return x_28; +lean_object* x_31; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +lean_ctor_set_tag(x_29, 1); +lean_ctor_set(x_29, 0, x_27); +return x_29; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_26); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_27); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } @@ -45890,23 +49503,6 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -45935,7 +49531,7 @@ x_14 = l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lam x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); -x_16 = l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4; +x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10; x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -45979,7 +49575,7 @@ x_25 = l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lam x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_24); -x_27 = l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4; +x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10; x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -49813,22 +53409,32 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Lean_Elab_Term_elabMutualDef(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_10; +lean_object* x_11; +x_11 = l_Lean_Elab_Term_elabMutualDef(x_3, x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; lean_object* x_7; -x_6 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMutualDef___lambda__1), 9, 1); -lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Elab_Command_runTermElabM___rarg(x_6, x_3, x_4, x_5); -return x_7; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = l_Lean_Elab_Command_getScope___rarg(x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_ctor_get(x_7, 6); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMutualDef___lambda__1), 10, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_1); +x_11 = l_Lean_Elab_Command_runTermElabM___rarg(x_10, x_3, x_4, x_8); +return x_11; } } static lean_object* _init_l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__1() { @@ -50371,7 +53977,7 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -50381,27 +53987,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1; x_2 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4() { _start: { lean_object* x_1; @@ -50409,17 +54015,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6() { _start: { lean_object* x_1; @@ -50427,47 +54033,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8; x_2 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11() { _start: { lean_object* x_1; @@ -50475,33 +54081,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12; -x_2 = lean_unsigned_to_nat(15717u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12; +x_2 = lean_unsigned_to_nat(16546u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -50906,18 +54512,61 @@ l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab lean_mark_persistent(l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__1); l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__2 = _init_l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__2(); lean_mark_persistent(l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__2); -l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1(); +l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1 = _init_l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_instantiateMVarsProfiling___closed__1); +l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__1); +l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642____closed__6); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_5642_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_Elab_Term_deprecated_oldSectionVars = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_Elab_Term_deprecated_oldSectionVars); +lean_dec_ref(res); +}l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__2___closed__3); -l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1(); -lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1); -l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2(); -lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__1); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__2); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__3); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__4); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__5); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__6); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__7); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__8); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__9); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___closed__10); +l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1 = _init_l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__1); +l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2 = _init_l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__10___lambda__5___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__15___lambda__1___closed__3); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__2(); @@ -51054,26 +54703,22 @@ l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__2(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__2); -l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1(); -lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__1); -l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2(); -lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__2); -l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3(); -lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__8___closed__3); -l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1 = _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1(); -lean_mark_persistent(l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__1); -l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2 = _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2(); -lean_mark_persistent(l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__6___closed__2); +l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1(); +lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__1); +l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2(); +lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__2); +l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3 = _init_l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3(); +lean_mark_persistent(l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__9___closed__3); +l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1 = _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1(); +lean_mark_persistent(l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__1); +l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2 = _init_l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2(); +lean_mark_persistent(l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__7___closed__2); l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___closed__1 = _init_l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___closed__1(); lean_mark_persistent(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___closed__1); l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__1 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__1); l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__2 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__2(); lean_mark_persistent(l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__2); -l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3); -l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4); l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1(); lean_mark_persistent(l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1); l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__2 = _init_l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__2(); @@ -51122,33 +54767,33 @@ l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__4 = _init_l_Lean_Elab_Co lean_mark_persistent(l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__4); l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5 = _init_l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5(); lean_mark_persistent(l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717____closed__13); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_15717_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546____closed__13); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16546_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c index 40c7afb535f3..200f1c0cd258 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c @@ -37,7 +37,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__35___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__13; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_evalOmega___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__4; lean_object* l_Lean_mkNatLit(lean_object*); @@ -190,6 +189,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__35___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_evalOmega(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525_(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__33___closed__1; @@ -292,6 +292,7 @@ static lean_object* l_List_map___at_Lean_Elab_Tactic_Omega_formatErrorMessage_pr LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__29; @@ -338,7 +339,6 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Tactic_Omega_formatErrorMessage_inScope___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__19; LEAN_EXPORT lean_object* l_List_map___at_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyAtoms___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539_(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__35___closed__3; @@ -349,7 +349,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__8; lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_mentioned(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__40; @@ -467,9 +467,9 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___clo LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyConstraint___closed__4; lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__46; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__19___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -519,7 +519,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2(lean LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Omega_LinearCombo_isAtom(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__19___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -600,7 +599,6 @@ static lean_object* l_List_map___at_Lean_Elab_Tactic_Omega_formatErrorMessage_pr static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__42; lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3___boxed(lean_object**); -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__8___closed__1; @@ -624,6 +622,7 @@ lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean static lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082___closed__2; LEAN_EXPORT lean_object* l_List_map___at_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyCoeffs___spec__1___boxed(lean_object*, lean_object*); uint8_t l_List_decidableBAll___rarg(lean_object*, lean_object*); @@ -801,7 +800,6 @@ lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMe lean_object* l_Int_toNat(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyAtoms___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__23; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -889,7 +887,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__12; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__3___closed__2; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyConstraints(lean_object*, lean_object*); @@ -898,6 +895,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__19; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__8___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -972,10 +970,11 @@ lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); lean_object* l_Lean_Omega_LinearCombo_mul(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__8; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__3; +static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Tactic_Omega_omegaTactic___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_int_neg(lean_object*); @@ -6348,97 +6347,109 @@ if (lean_obj_tag(x_16) == 0) lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); lean_dec(x_16); -x_19 = lean_st_ref_take(x_2, x_18); -x_20 = !lean_is_exclusive(x_19); +x_20 = !lean_is_exclusive(x_17); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -x_23 = lean_ctor_get(x_17, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_17, 1); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +x_26 = lean_st_ref_take(x_2, x_19); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); +lean_closure_set(x_30, 0, x_28); +x_31 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_31, 0, x_30); lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec(x_24); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); +lean_closure_set(x_32, 0, x_24); +x_33 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; +x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); lean_inc(x_21); -x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); -lean_closure_set(x_26, 0, x_21); -x_27 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_27, 0, x_26); -x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); -lean_closure_set(x_28, 0, x_25); -x_29 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); -lean_closure_set(x_29, 0, x_27); -lean_closure_set(x_29, 1, x_28); -x_30 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; -x_31 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); -lean_closure_set(x_31, 0, x_29); -lean_closure_set(x_31, 1, x_30); -lean_ctor_set(x_19, 1, x_31); -lean_ctor_set(x_19, 0, x_23); -x_32 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_21, x_1, x_19); -x_33 = lean_st_ref_set(x_2, x_32, x_22); +lean_ctor_set(x_18, 1, x_35); +lean_ctor_set(x_18, 0, x_21); +x_36 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_28, x_1, x_18); +x_37 = lean_st_ref_set(x_2, x_36, x_29); lean_dec(x_2); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_35; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -lean_ctor_set(x_33, 0, x_17); -return x_33; +lean_object* x_39; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +lean_ctor_set(x_17, 1, x_25); +lean_ctor_set(x_17, 0, x_24); +lean_ctor_set(x_26, 1, x_17); +lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_37, 0, x_26); +return x_37; } else { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_17); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +lean_ctor_set(x_17, 1, x_25); +lean_ctor_set(x_17, 0, x_24); +lean_ctor_set(x_26, 1, x_17); +lean_ctor_set(x_26, 0, x_21); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_26); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_38 = lean_ctor_get(x_19, 0); -x_39 = lean_ctor_get(x_19, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_19); -x_40 = lean_ctor_get(x_17, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_17, 1); -lean_inc(x_41); -x_42 = lean_ctor_get(x_41, 0); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); lean_inc(x_42); -lean_dec(x_41); -lean_inc(x_38); -x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); -lean_closure_set(x_43, 0, x_38); -x_44 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_44, 0, x_43); -x_45 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); -lean_closure_set(x_45, 0, x_42); -x_46 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); -lean_closure_set(x_46, 0, x_44); -lean_closure_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; -x_48 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); -lean_closure_set(x_48, 0, x_46); -lean_closure_set(x_48, 1, x_47); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_40); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_38, x_1, x_49); -x_51 = lean_st_ref_set(x_2, x_50, x_39); +lean_dec(x_26); +lean_inc(x_42); +x_44 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); +lean_closure_set(x_44, 0, x_42); +x_45 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_45, 0, x_44); +lean_inc(x_24); +x_46 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); +lean_closure_set(x_46, 0, x_24); +x_47 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_47, 0, x_45); +lean_closure_set(x_47, 1, x_46); +x_48 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; +x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_49, 0, x_47); +lean_closure_set(x_49, 1, x_48); +lean_inc(x_21); +lean_ctor_set(x_18, 1, x_49); +lean_ctor_set(x_18, 0, x_21); +x_50 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_42, x_1, x_18); +x_51 = lean_st_ref_set(x_2, x_50, x_43); lean_dec(x_2); x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); @@ -6450,69 +6461,233 @@ if (lean_is_exclusive(x_51)) { lean_dec_ref(x_51); x_53 = lean_box(0); } +lean_ctor_set(x_17, 1, x_25); +lean_ctor_set(x_17, 0, x_24); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_21); +lean_ctor_set(x_54, 1, x_17); if (lean_is_scalar(x_53)) { - x_54 = lean_alloc_ctor(0, 2, 0); + x_55 = lean_alloc_ctor(0, 2, 0); } else { - x_54 = x_53; + x_55 = x_53; } -lean_ctor_set(x_54, 0, x_17); -lean_ctor_set(x_54, 1, x_52); -return x_54; +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; } } else { -uint8_t x_55; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_56 = lean_ctor_get(x_18, 0); +x_57 = lean_ctor_get(x_18, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_18); +x_58 = lean_st_ref_take(x_2, x_19); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_61 = x_58; +} else { + lean_dec_ref(x_58); + x_61 = lean_box(0); +} +lean_inc(x_59); +x_62 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); +lean_closure_set(x_62, 0, x_59); +x_63 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_63, 0, x_62); +lean_inc(x_56); +x_64 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); +lean_closure_set(x_64, 0, x_56); +x_65 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_65, 0, x_63); +lean_closure_set(x_65, 1, x_64); +x_66 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; +x_67 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_67, 0, x_65); +lean_closure_set(x_67, 1, x_66); +lean_inc(x_21); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_21); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_59, x_1, x_68); +x_70 = lean_st_ref_set(x_2, x_69, x_60); +lean_dec(x_2); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_72 = x_70; +} else { + lean_dec_ref(x_70); + x_72 = lean_box(0); +} +lean_ctor_set(x_17, 1, x_57); +lean_ctor_set(x_17, 0, x_56); +if (lean_is_scalar(x_61)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = x_61; +} +lean_ctor_set(x_73, 0, x_21); +lean_ctor_set(x_73, 1, x_17); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_72; +} +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_71); +return x_74; +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_75 = lean_ctor_get(x_17, 0); +lean_inc(x_75); +lean_dec(x_17); +x_76 = lean_ctor_get(x_18, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_18, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_78 = x_18; +} else { + lean_dec_ref(x_18); + x_78 = lean_box(0); +} +x_79 = lean_st_ref_take(x_2, x_19); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_82 = x_79; +} else { + lean_dec_ref(x_79); + x_82 = lean_box(0); +} +lean_inc(x_80); +x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); +lean_closure_set(x_83, 0, x_80); +x_84 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_84, 0, x_83); +lean_inc(x_76); +x_85 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); +lean_closure_set(x_85, 0, x_76); +x_86 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_86, 0, x_84); +lean_closure_set(x_86, 1, x_85); +x_87 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; +x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_88, 0, x_86); +lean_closure_set(x_88, 1, x_87); +lean_inc(x_75); +if (lean_is_scalar(x_78)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_78; +} +lean_ctor_set(x_89, 0, x_75); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_80, x_1, x_89); +x_91 = lean_st_ref_set(x_2, x_90, x_81); +lean_dec(x_2); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_93 = x_91; +} else { + lean_dec_ref(x_91); + x_93 = lean_box(0); +} +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_76); +lean_ctor_set(x_94, 1, x_77); +if (lean_is_scalar(x_82)) { + x_95 = lean_alloc_ctor(0, 2, 0); +} else { + x_95 = x_82; +} +lean_ctor_set(x_95, 0, x_75); +lean_ctor_set(x_95, 1, x_94); +if (lean_is_scalar(x_93)) { + x_96 = lean_alloc_ctor(0, 2, 0); +} else { + x_96 = x_93; +} +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_92); +return x_96; +} +} +else +{ +uint8_t x_97; lean_dec(x_2); lean_dec(x_1); -x_55 = !lean_is_exclusive(x_16); -if (x_55 == 0) +x_97 = !lean_is_exclusive(x_16); +if (x_97 == 0) { return x_16; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_16, 0); -x_57 = lean_ctor_get(x_16, 1); -lean_inc(x_57); -lean_inc(x_56); +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_16, 0); +x_99 = lean_ctor_get(x_16, 1); +lean_inc(x_99); +lean_inc(x_98); lean_dec(x_16); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -lean_object* x_59; uint8_t x_60; -x_59 = lean_ctor_get(x_15, 0); -lean_inc(x_59); +lean_object* x_101; uint8_t x_102; +x_101 = lean_ctor_get(x_15, 0); +lean_inc(x_101); lean_dec(x_15); -x_60 = !lean_is_exclusive(x_59); -if (x_60 == 0) +x_102 = !lean_is_exclusive(x_101); +if (x_102 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_61 = lean_ctor_get(x_59, 0); -x_62 = lean_ctor_get(x_59, 1); -x_63 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_64 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_unbox(x_65); -lean_dec(x_65); -if (x_66 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = lean_ctor_get(x_101, 0); +x_104 = lean_ctor_get(x_101, 1); +x_105 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_106 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -lean_free_object(x_59); +lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_free_object(x_101); lean_dec(x_1); -x_67 = lean_ctor_get(x_64, 1); -lean_inc(x_67); -lean_dec(x_64); -x_68 = lean_box(0); -x_69 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_62, x_61, x_68, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_67); +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +lean_dec(x_106); +x_110 = lean_box(0); +x_111 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_104, x_103, x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_109); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6521,34 +6696,34 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_69; +return x_111; } else { -uint8_t x_70; -x_70 = !lean_is_exclusive(x_64); -if (x_70 == 0) +uint8_t x_112; +x_112 = !lean_is_exclusive(x_106); +if (x_112 == 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; -x_71 = lean_ctor_get(x_64, 1); -x_72 = lean_ctor_get(x_64, 0); -lean_dec(x_72); -x_73 = l_Lean_MessageData_ofExpr(x_1); -x_74 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; -lean_ctor_set_tag(x_64, 7); -lean_ctor_set(x_64, 1, x_73); -lean_ctor_set(x_64, 0, x_74); -x_75 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -lean_ctor_set_tag(x_59, 7); -lean_ctor_set(x_59, 1, x_75); -lean_ctor_set(x_59, 0, x_64); -x_76 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_63, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_71); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_62, x_61, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_78); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_113 = lean_ctor_get(x_106, 1); +x_114 = lean_ctor_get(x_106, 0); +lean_dec(x_114); +x_115 = l_Lean_MessageData_ofExpr(x_1); +x_116 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; +lean_ctor_set_tag(x_106, 7); +lean_ctor_set(x_106, 1, x_115); +lean_ctor_set(x_106, 0, x_116); +x_117 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_ctor_set_tag(x_101, 7); +lean_ctor_set(x_101, 1, x_117); +lean_ctor_set(x_101, 0, x_106); +x_118 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_105, x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_113); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_104, x_103, x_119, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_120); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6557,31 +6732,31 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_77); -return x_79; +lean_dec(x_119); +return x_121; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_80 = lean_ctor_get(x_64, 1); -lean_inc(x_80); -lean_dec(x_64); -x_81 = l_Lean_MessageData_ofExpr(x_1); -x_82 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -lean_ctor_set_tag(x_59, 7); -lean_ctor_set(x_59, 1, x_84); -lean_ctor_set(x_59, 0, x_83); -x_85 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_63, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_80); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -lean_dec(x_85); -x_88 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_62, x_61, x_86, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_87); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_122 = lean_ctor_get(x_106, 1); +lean_inc(x_122); +lean_dec(x_106); +x_123 = l_Lean_MessageData_ofExpr(x_1); +x_124 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; +x_125 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_123); +x_126 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_ctor_set_tag(x_101, 7); +lean_ctor_set(x_101, 1, x_126); +lean_ctor_set(x_101, 0, x_125); +x_127 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_105, x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_122); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_104, x_103, x_128, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_129); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6590,34 +6765,34 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_86); -return x_88; +lean_dec(x_128); +return x_130; } } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; -x_89 = lean_ctor_get(x_59, 0); -x_90 = lean_ctor_get(x_59, 1); -lean_inc(x_90); -lean_inc(x_89); -lean_dec(x_59); -x_91 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_92 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_91, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_unbox(x_93); -lean_dec(x_93); -if (x_94 == 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; +x_131 = lean_ctor_get(x_101, 0); +x_132 = lean_ctor_get(x_101, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_101); +x_133 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_134 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_unbox(x_135); +lean_dec(x_135); +if (x_136 == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_dec(x_1); -x_95 = lean_ctor_get(x_92, 1); -lean_inc(x_95); -lean_dec(x_92); -x_96 = lean_box(0); -x_97 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_90, x_89, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_95); +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +lean_dec(x_134); +x_138 = lean_box(0); +x_139 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_132, x_131, x_138, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_137); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6626,42 +6801,42 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_97; +return x_139; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_98 = lean_ctor_get(x_92, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_99 = x_92; +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_140 = lean_ctor_get(x_134, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_141 = x_134; } else { - lean_dec_ref(x_92); - x_99 = lean_box(0); + lean_dec_ref(x_134); + x_141 = lean_box(0); } -x_100 = l_Lean_MessageData_ofExpr(x_1); -x_101 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; -if (lean_is_scalar(x_99)) { - x_102 = lean_alloc_ctor(7, 2, 0); +x_142 = l_Lean_MessageData_ofExpr(x_1); +x_143 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; +if (lean_is_scalar(x_141)) { + x_144 = lean_alloc_ctor(7, 2, 0); } else { - x_102 = x_99; - lean_ctor_set_tag(x_102, 7); + x_144 = x_141; + lean_ctor_set_tag(x_144, 7); } -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_100); -x_103 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_91, x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_98); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_90, x_89, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_107); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_142); +x_145 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_146 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_133, x_146, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_140); +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_132, x_131, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_149); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6670,8 +6845,8 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_106); -return x_108; +lean_dec(x_148); +return x_150; } } } @@ -13213,162 +13388,6 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_1); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_1, 1); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_1, 0); -x_16 = lean_ctor_get(x_13, 0); -x_17 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -x_18 = l_Lean_Omega_LinearCombo_isAtom(x_15); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_1); -x_20 = 1; -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_11); -return x_23; -} -else -{ -lean_object* x_24; lean_object* x_25; -lean_free_object(x_13); -lean_dec(x_17); -lean_dec(x_16); -lean_free_object(x_1); -lean_dec(x_15); -x_24 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__8; -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_11); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_1, 0); -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_13); -lean_inc(x_26); -x_29 = l_Lean_Omega_LinearCombo_isAtom(x_26); -if (x_29 == 0) -{ -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_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_27); -lean_ctor_set(x_30, 1, x_28); -lean_ctor_set(x_1, 1, x_30); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_1); -x_32 = 1; -x_33 = lean_box(x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_11); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; -lean_dec(x_28); -lean_dec(x_27); -lean_free_object(x_1); -lean_dec(x_26); -x_36 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__8; -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_11); -return x_37; -} -} -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_38 = lean_ctor_get(x_1, 1); -x_39 = lean_ctor_get(x_1, 0); -lean_inc(x_38); -lean_inc(x_39); -lean_dec(x_1); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -lean_inc(x_39); -x_43 = l_Lean_Omega_LinearCombo_isAtom(x_39); -if (x_43 == 0) -{ -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; -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_41); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_39); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_46, 0, x_45); -x_47 = 1; -x_48 = lean_box(x_47); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_46); -lean_ctor_set(x_49, 1, x_48); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_11); -return x_50; -} -else -{ -lean_object* x_51; lean_object* x_52; -lean_dec(x_42); -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__8; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_11); -return x_52; -} -} -} -} static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__1() { _start: { @@ -13809,39 +13828,61 @@ static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCa _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1___boxed), 11, 0); +x_1 = lean_mk_string_unchecked("ofNat_mul_nonneg", 16, 16); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__48() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__1; +x_2 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__3; +x_3 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_4 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__47; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__48; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_unchecked("ofNat_add", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__48; +x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__51; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__51() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__53() { _start: { lean_object* x_1; @@ -13849,27 +13890,27 @@ x_1 = lean_mk_string_unchecked("ofNat_succ", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__51; +x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__53; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__53() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52; +x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13879,11 +13920,11 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; +x_1 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56; x_2 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__10___closed__30; x_3 = l_Lean_Expr_const___override(x_1, x_2); return x_3; @@ -13928,7 +13969,7 @@ lean_dec(x_15); x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55; +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57; x_22 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; x_23 = l_Lean_mkApp3(x_21, x_22, x_2, x_20); lean_inc(x_12); @@ -15800,118 +15841,173 @@ return x_375; } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; +lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; x_376 = lean_unsigned_to_nat(4u); x_377 = lean_array_fget(x_40, x_376); x_378 = lean_unsigned_to_nat(5u); x_379 = lean_array_fget(x_40, x_378); lean_dec(x_40); x_380 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__46; +lean_inc(x_379); +lean_inc(x_377); x_381 = l_Lean_mkAppB(x_380, x_377, x_379); -lean_inc(x_1); -x_382 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___boxed), 12, 2); -lean_closure_set(x_382, 0, x_1); -lean_closure_set(x_382, 1, x_381); -x_383 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__47; -x_384 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); -lean_closure_set(x_384, 0, x_382); -lean_closure_set(x_384, 1, x_383); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_385 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_384, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_385) == 0) +x_382 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_381, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_382) == 0) { -lean_object* x_386; -x_386 = lean_ctor_get(x_385, 0); -lean_inc(x_386); -if (lean_obj_tag(x_386) == 0) +lean_object* x_383; lean_object* x_384; uint8_t x_385; +x_383 = lean_ctor_get(x_382, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_383, 1); +lean_inc(x_384); +x_385 = !lean_is_exclusive(x_382); +if (x_385 == 0) { -lean_object* x_387; lean_object* x_388; -x_387 = lean_ctor_get(x_385, 1); -lean_inc(x_387); -lean_dec(x_385); -x_388 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_387); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_388; +lean_object* x_386; uint8_t x_387; +x_386 = lean_ctor_get(x_382, 0); +lean_dec(x_386); +x_387 = !lean_is_exclusive(x_383); +if (x_387 == 0) +{ +lean_object* x_388; uint8_t x_389; +x_388 = lean_ctor_get(x_383, 1); +lean_dec(x_388); +x_389 = !lean_is_exclusive(x_384); +if (x_389 == 0) +{ +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; +x_390 = lean_ctor_get(x_384, 1); +x_391 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_392 = l_Lean_mkAppB(x_391, x_377, x_379); +x_393 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_390, x_392); +lean_ctor_set(x_384, 1, x_393); +return x_382; } else { -uint8_t x_389; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_389 = !lean_is_exclusive(x_385); -if (x_389 == 0) +lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; +x_394 = lean_ctor_get(x_384, 0); +x_395 = lean_ctor_get(x_384, 1); +lean_inc(x_395); +lean_inc(x_394); +lean_dec(x_384); +x_396 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_397 = l_Lean_mkAppB(x_396, x_377, x_379); +x_398 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_395, x_397); +x_399 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_399, 0, x_394); +lean_ctor_set(x_399, 1, x_398); +lean_ctor_set(x_383, 1, x_399); +return x_382; +} +} +else { -lean_object* x_390; lean_object* x_391; -x_390 = lean_ctor_get(x_385, 0); -lean_dec(x_390); -x_391 = lean_ctor_get(x_386, 0); -lean_inc(x_391); -lean_dec(x_386); -lean_ctor_set(x_385, 0, x_391); -return x_385; +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; +x_400 = lean_ctor_get(x_383, 0); +lean_inc(x_400); +lean_dec(x_383); +x_401 = lean_ctor_get(x_384, 0); +lean_inc(x_401); +x_402 = lean_ctor_get(x_384, 1); +lean_inc(x_402); +if (lean_is_exclusive(x_384)) { + lean_ctor_release(x_384, 0); + lean_ctor_release(x_384, 1); + x_403 = x_384; +} else { + lean_dec_ref(x_384); + x_403 = lean_box(0); +} +x_404 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_405 = l_Lean_mkAppB(x_404, x_377, x_379); +x_406 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_402, x_405); +if (lean_is_scalar(x_403)) { + x_407 = lean_alloc_ctor(0, 2, 0); +} else { + x_407 = x_403; +} +lean_ctor_set(x_407, 0, x_401); +lean_ctor_set(x_407, 1, x_406); +x_408 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_408, 0, x_400); +lean_ctor_set(x_408, 1, x_407); +lean_ctor_set(x_382, 0, x_408); +return x_382; +} } else { -lean_object* x_392; lean_object* x_393; lean_object* x_394; -x_392 = lean_ctor_get(x_385, 1); -lean_inc(x_392); -lean_dec(x_385); -x_393 = lean_ctor_get(x_386, 0); -lean_inc(x_393); -lean_dec(x_386); -x_394 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_394, 0, x_393); -lean_ctor_set(x_394, 1, x_392); -return x_394; +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; +x_409 = lean_ctor_get(x_382, 1); +lean_inc(x_409); +lean_dec(x_382); +x_410 = lean_ctor_get(x_383, 0); +lean_inc(x_410); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_411 = x_383; +} else { + lean_dec_ref(x_383); + x_411 = lean_box(0); } +x_412 = lean_ctor_get(x_384, 0); +lean_inc(x_412); +x_413 = lean_ctor_get(x_384, 1); +lean_inc(x_413); +if (lean_is_exclusive(x_384)) { + lean_ctor_release(x_384, 0); + lean_ctor_release(x_384, 1); + x_414 = x_384; +} else { + lean_dec_ref(x_384); + x_414 = lean_box(0); +} +x_415 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_416 = l_Lean_mkAppB(x_415, x_377, x_379); +x_417 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_413, x_416); +if (lean_is_scalar(x_414)) { + x_418 = lean_alloc_ctor(0, 2, 0); +} else { + x_418 = x_414; +} +lean_ctor_set(x_418, 0, x_412); +lean_ctor_set(x_418, 1, x_417); +if (lean_is_scalar(x_411)) { + x_419 = lean_alloc_ctor(0, 2, 0); +} else { + x_419 = x_411; +} +lean_ctor_set(x_419, 0, x_410); +lean_ctor_set(x_419, 1, x_418); +x_420 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_420, 0, x_419); +lean_ctor_set(x_420, 1, x_409); +return x_420; } } else { -uint8_t x_395; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_395 = !lean_is_exclusive(x_385); -if (x_395 == 0) +uint8_t x_421; +lean_dec(x_379); +lean_dec(x_377); +x_421 = !lean_is_exclusive(x_382); +if (x_421 == 0) { -return x_385; +return x_382; } else { -lean_object* x_396; lean_object* x_397; lean_object* x_398; -x_396 = lean_ctor_get(x_385, 0); -x_397 = lean_ctor_get(x_385, 1); -lean_inc(x_397); -lean_inc(x_396); -lean_dec(x_385); -x_398 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_398, 0, x_396); -lean_ctor_set(x_398, 1, x_397); -return x_398; +lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_422 = lean_ctor_get(x_382, 0); +x_423 = lean_ctor_get(x_382, 1); +lean_inc(x_423); +lean_inc(x_422); +lean_dec(x_382); +x_424 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_424, 0, x_422); +lean_ctor_set(x_424, 1, x_423); +return x_424; } } } @@ -15920,151 +16016,151 @@ return x_398; } else { -lean_object* x_399; uint8_t x_400; +lean_object* x_425; uint8_t x_426; lean_dec(x_42); lean_dec(x_2); -x_399 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; -x_400 = lean_string_dec_eq(x_41, x_399); +x_425 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; +x_426 = lean_string_dec_eq(x_41, x_425); lean_dec(x_41); -if (x_400 == 0) +if (x_426 == 0) { -lean_object* x_401; +lean_object* x_427; lean_dec(x_40); -x_401 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_427 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_401; +return x_427; } else { -lean_object* x_402; lean_object* x_403; uint8_t x_404; -x_402 = lean_array_get_size(x_40); -x_403 = lean_unsigned_to_nat(6u); -x_404 = lean_nat_dec_eq(x_402, x_403); -lean_dec(x_402); -if (x_404 == 0) +lean_object* x_428; lean_object* x_429; uint8_t x_430; +x_428 = lean_array_get_size(x_40); +x_429 = lean_unsigned_to_nat(6u); +x_430 = lean_nat_dec_eq(x_428, x_429); +lean_dec(x_428); +if (x_430 == 0) { -lean_object* x_405; +lean_object* x_431; lean_dec(x_40); -x_405 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_431 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_405; +return x_431; } else { -lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; -x_406 = lean_unsigned_to_nat(4u); -x_407 = lean_array_fget(x_40, x_406); -x_408 = lean_unsigned_to_nat(5u); -x_409 = lean_array_fget(x_40, x_408); +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; +x_432 = lean_unsigned_to_nat(4u); +x_433 = lean_array_fget(x_40, x_432); +x_434 = lean_unsigned_to_nat(5u); +x_435 = lean_array_fget(x_40, x_434); lean_dec(x_40); -x_410 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50; -x_411 = l_Lean_mkAppB(x_410, x_407, x_409); -x_412 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_411, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_412; +x_436 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52; +x_437 = l_Lean_mkAppB(x_436, x_433, x_435); +x_438 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_437, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_438; } } } } else { -lean_object* x_413; uint8_t x_414; +lean_object* x_439; uint8_t x_440; lean_dec(x_42); lean_dec(x_2); -x_413 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__5; -x_414 = lean_string_dec_eq(x_41, x_413); +x_439 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__5; +x_440 = lean_string_dec_eq(x_41, x_439); lean_dec(x_41); -if (x_414 == 0) +if (x_440 == 0) { -lean_object* x_415; +lean_object* x_441; lean_dec(x_40); -x_415 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_441 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_415; +return x_441; } else { -lean_object* x_416; lean_object* x_417; uint8_t x_418; -x_416 = lean_array_get_size(x_40); -x_417 = lean_unsigned_to_nat(1u); -x_418 = lean_nat_dec_eq(x_416, x_417); -lean_dec(x_416); -if (x_418 == 0) +lean_object* x_442; lean_object* x_443; uint8_t x_444; +x_442 = lean_array_get_size(x_40); +x_443 = lean_unsigned_to_nat(1u); +x_444 = lean_nat_dec_eq(x_442, x_443); +lean_dec(x_442); +if (x_444 == 0) { -lean_object* x_419; +lean_object* x_445; lean_dec(x_40); -x_419 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_445 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_419; +return x_445; } else { -lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; -x_420 = lean_unsigned_to_nat(0u); -x_421 = lean_array_fget(x_40, x_420); +lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; +x_446 = lean_unsigned_to_nat(0u); +x_447 = lean_array_fget(x_40, x_446); lean_dec(x_40); -x_422 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__53; -x_423 = l_Lean_Expr_app___override(x_422, x_421); -x_424 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_423, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_424; +x_448 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55; +x_449 = l_Lean_Expr_app___override(x_448, x_447); +x_450 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_449, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_450; } } } } else { -lean_object* x_425; +lean_object* x_451; lean_dec(x_39); lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_2); -x_425 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_451 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_425; +return x_451; } } else { -lean_object* x_426; +lean_object* x_452; lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_dec(x_2); -x_426 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_452 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_426; +return x_452; } } else { -lean_object* x_427; +lean_object* x_453; lean_dec(x_37); lean_dec(x_36); lean_dec(x_2); -x_427 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_453 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_427; +return x_453; } } } @@ -16251,7 +16347,7 @@ lean_inc(x_21); lean_dec(x_17); x_22 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__14; x_23 = l_Lean_mkAppB(x_22, x_3, x_21); -x_24 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55; +x_24 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57; x_25 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; x_26 = l_Lean_mkApp3(x_24, x_25, x_2, x_23); lean_inc(x_13); @@ -17204,24 +17300,6 @@ x_13 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl(x_1, x_2, x_3, x_4, x_12, x_6, return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_5); -lean_dec(x_5); -x_13 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___lambda__1(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_13; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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: { @@ -33644,7 +33722,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1() { _start: { lean_object* x_1; @@ -33652,17 +33730,17 @@ x_1 = lean_mk_string_unchecked("bv_toNat", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1; +x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3() { _start: { lean_object* x_1; @@ -33670,7 +33748,7 @@ x_1 = lean_mk_string_unchecked("bvOmegaSimpExtension", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -33678,12 +33756,12 @@ x_1 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend x_2 = l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega__1___closed__1; x_3 = l_Lean_Elab_Tactic_Omega_evalOmega___closed__2; x_4 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__3; -x_5 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3; +x_5 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5() { _start: { lean_object* x_1; @@ -33691,13 +33769,13 @@ x_1 = lean_mk_string_unchecked("simp lemmas converting `BitVec` goals to `Nat` g return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2; -x_3 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5; -x_4 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4; +x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2; +x_3 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5; +x_4 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4; x_5 = l_Lean_Meta_registerSimpAttr(x_2, x_3, x_4, x_1); return x_5; } @@ -34185,6 +34263,10 @@ l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54 = _init_l_ lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54); l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__55); +l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__56); +l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__57); l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__1 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__1); l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__2 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__2(); @@ -34717,17 +34799,17 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange__ if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__1); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__2); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__3); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__4); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539____closed__5); -if (builtin) {res = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15539_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__1); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__2); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__3); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__4); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525____closed__5); +if (builtin) {res = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15525_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Tactic_Omega_bvOmegaSimpExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_bvOmegaSimpExtension); diff --git a/stage0/stdlib/Lean/Environment.c b/stage0/stdlib/Lean/Environment.c index 0575544247f3..ba50bc500b7a 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -43,17 +43,20 @@ LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_quickLt___boxed(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_evalConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_ImportStateM_run___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3302_; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); lean_object* lean_add_decl(lean_object*, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2; lean_object* lean_uint32_to_nat(uint32_t); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_equivInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1; @@ -80,13 +83,14 @@ static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1 LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_regions___default; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3; LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__3; LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_finalizeImport___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443_(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_setImportedEntries(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -96,6 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Environme LEAN_EXPORT uint8_t l_Lean_EnvironmentHeader_quotInit___default; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2607_; @@ -113,6 +118,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries(lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__3(lean_object*); @@ -142,7 +148,6 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg__ static lean_object* l_Lean_Kernel_resetDiag___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_importModules___lambda__2___closed__1; @@ -155,13 +160,12 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_num_attributes(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_moduleNames___default; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__11; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Environment_find_x3f___spec__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6572_(lean_object*); LEAN_EXPORT lean_object* lean_environment_find(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -178,6 +182,7 @@ LEAN_EXPORT lean_object* l_Lean_getNumBuiltinAttributes___boxed(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__21; static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* lean_kernel_get_diag(lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*, lean_object*); @@ -196,15 +201,17 @@ LEAN_EXPORT lean_object* l_Lean_Environment_addExtraName(lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Environment_isNamespace___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3; LEAN_EXPORT lean_object* l_Lean_EnvExtensionEntrySpec; LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_importModules___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_System_FilePath_pathExists(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__5(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5; LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Environment_addExtraName___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getNamespaceSet(lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,20 +238,18 @@ lean_object* l_Lean_initializing(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__22; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6___closed__4; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6476_(lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6___lambda__1___closed__3; lean_object* l_IO_println___at_Lean_instEval___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__6; LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtensionDescr_toArrayFn___default(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347_(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__3___boxed(lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); @@ -254,7 +259,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___s LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_mkExtNameMap___spec__3___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportState_moduleNames___default; uint8_t l_Lean_HashSetImp_contains___at_Lean_NameHashSet_contains___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_Diagnostics_isEnabled___boxed(lean_object*); @@ -293,7 +297,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeReg lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___boxed(lean_object*, lean_object*); static lean_object* l_Lean_mkMapDeclarationExtension___rarg___closed__2; @@ -316,7 +319,6 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l_Lean_instModuleIdxToString; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6; static lean_object* l_Lean_mkEmptyEnvironment___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_addAux(lean_object*, lean_object*); @@ -326,6 +328,7 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_finalizeImport___spec LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_getTrustLevel___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportState_regions___default; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__10; @@ -336,7 +339,6 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___bo LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__3; @@ -372,6 +374,7 @@ lean_object* lean_list_to_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__24; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__7; @@ -388,13 +391,11 @@ static lean_object* l_Lean_TagDeclarationExtension_tag___closed__3; static lean_object* l_Lean_instToStringImport___closed__1; LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_mkExtNameMap___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5; LEAN_EXPORT lean_object* l_Lean_mkDefinitionValInferrringUnsafe(lean_object*); LEAN_EXPORT uint32_t lean_environment_trust_level(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___boxed(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____boxed(lean_object*, lean_object*); @@ -423,8 +424,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isName LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__27; LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instToStringImport___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_Environment_isNamespaceName(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParams(lean_object*, lean_object*, lean_object*); @@ -434,6 +437,7 @@ static lean_object* l_Lean_TagDeclarationExtension_tag___closed__2; LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedImport; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___closed__2; @@ -463,8 +467,8 @@ LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environ LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2(lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__1; @@ -515,12 +519,11 @@ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___clos LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object*); LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object*); lean_object* lean_string_length(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__17; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__14; static lean_object* l_Lean_instReprImport___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object*, lean_object*, lean_object*); uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_433____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3458_; @@ -532,7 +535,6 @@ LEAN_EXPORT lean_object* l_Lean_instReprImport; static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -559,7 +561,7 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Environment_getMod LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__4(lean_object*, size_t, lean_object*); -lean_object* l_Lean_findOLean(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_findOLean(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Import_runtimeOnly___default; LEAN_EXPORT lean_object* l_Lean_Environment_isConstructor___boxed(lean_object*, lean_object*); @@ -581,7 +583,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spe LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object*, lean_object*); extern lean_object* l_Id_instMonad; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -638,6 +640,7 @@ lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_imports(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -679,9 +682,9 @@ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg___closed__1; static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__9; size_t lean_array_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__5; lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_CompactedRegion_isMemoryMapped___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__1(lean_object*, lean_object*, lean_object*); @@ -699,7 +702,9 @@ size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3f___spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__9(lean_object*, size_t, size_t, lean_object*); @@ -720,7 +725,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___ static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__19; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkExtNameMap___spec__1___boxed(lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__2; @@ -747,6 +753,7 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState(lean_objec lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_moduleData___default; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_import_modules(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); @@ -778,10 +785,10 @@ LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object*, lean_object lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__4; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2; static lean_object* l___auto____x40_Lean_Environment___hyg_2607____closed__17; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object*); @@ -797,7 +804,6 @@ lean_object* l_List_length___rarg(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__14; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6___lambda__1___closed__1; lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; @@ -11509,423 +11515,434 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_st_ref_take(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = !lean_is_exclusive(x_6); -if (x_8 == 0) +lean_object* x_6; +x_6 = lean_read_module_data(x_1, x_5); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_9 = lean_ctor_get(x_6, 0); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 1); lean_inc(x_10); -x_11 = l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_9, x_10); -lean_ctor_set(x_6, 0, x_11); -x_12 = lean_st_ref_set(x_3, x_6, x_7); +lean_dec(x_7); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +x_12 = l_Lean_importModulesCore(x_11, x_4, x_8); +lean_dec(x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = 1; -lean_inc(x_10); -x_15 = l_Lean_findOLean(x_10, x_14, x_13); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 0); +x_14 = lean_st_ref_take(x_4, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_read_module_data(x_16, x_17); -lean_dec(x_16); -if (lean_obj_tag(x_18) == 0) +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -x_24 = l_Lean_importModulesCore(x_23, x_3, x_20); -lean_dec(x_23); -if (lean_obj_tag(x_24) == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_18 = lean_ctor_get(x_15, 1); +x_19 = lean_ctor_get(x_15, 2); +x_20 = lean_ctor_get(x_15, 3); +x_21 = lean_array_push(x_18, x_2); +x_22 = lean_array_push(x_19, x_9); +x_23 = lean_array_push(x_20, x_10); +lean_ctor_set(x_15, 3, x_23); +lean_ctor_set(x_15, 2, x_22); +lean_ctor_set(x_15, 1, x_21); +x_24 = lean_st_ref_set(x_4, x_15, x_16); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_st_ref_take(x_3, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); lean_dec(x_26); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_30 = lean_ctor_get(x_27, 1); -x_31 = lean_ctor_get(x_27, 2); -x_32 = lean_ctor_get(x_27, 3); -x_33 = lean_array_push(x_30, x_10); -x_34 = lean_array_push(x_31, x_21); -x_35 = lean_array_push(x_32, x_22); -lean_ctor_set(x_27, 3, x_35); -lean_ctor_set(x_27, 2, x_34); -lean_ctor_set(x_27, 1, x_33); -x_36 = lean_st_ref_set(x_3, x_27, x_28); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_36, 0); -lean_dec(x_38); -x_39 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -lean_ctor_set(x_36, 0, x_39); -return x_36; +x_27 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; +lean_ctor_set(x_24, 0, x_27); +return x_24; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 1); -lean_inc(x_40); -lean_dec(x_36); -x_41 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +x_29 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_43 = lean_ctor_get(x_27, 0); -x_44 = lean_ctor_get(x_27, 1); -x_45 = lean_ctor_get(x_27, 2); -x_46 = lean_ctor_get(x_27, 3); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_27); -x_47 = lean_array_push(x_44, x_10); -x_48 = lean_array_push(x_45, x_21); -x_49 = lean_array_push(x_46, x_22); -x_50 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_47); -lean_ctor_set(x_50, 2, x_48); -lean_ctor_set(x_50, 3, x_49); -x_51 = lean_st_ref_set(x_3, x_50, x_28); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_53 = x_51; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); +x_33 = lean_ctor_get(x_15, 2); +x_34 = lean_ctor_get(x_15, 3); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_15); +x_35 = lean_array_push(x_32, x_2); +x_36 = lean_array_push(x_33, x_9); +x_37 = lean_array_push(x_34, x_10); +x_38 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_35); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_37); +x_39 = lean_st_ref_set(x_4, x_38, x_16); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } else { - lean_dec_ref(x_51); - x_53 = lean_box(0); + lean_dec_ref(x_39); + x_41 = lean_box(0); } -x_54 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -if (lean_is_scalar(x_53)) { - x_55 = lean_alloc_ctor(0, 2, 0); +x_42 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_53; + x_43 = x_41; } -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_52); -return x_55; +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } else { -uint8_t x_56; -lean_dec(x_22); -lean_dec(x_21); +uint8_t x_44; lean_dec(x_10); -x_56 = !lean_is_exclusive(x_24); -if (x_56 == 0) +lean_dec(x_9); +lean_dec(x_2); +x_44 = !lean_is_exclusive(x_12); +if (x_44 == 0) { -return x_24; +return x_12; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_24, 0); -x_58 = lean_ctor_get(x_24, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_24); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -uint8_t x_60; -lean_dec(x_10); -x_60 = !lean_is_exclusive(x_18); -if (x_60 == 0) +uint8_t x_48; +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_6); +if (x_48 == 0) { -return x_18; +return x_6; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_18, 0); -x_62 = lean_ctor_get(x_18, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_18); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_6, 0); +x_50 = lean_ctor_get(x_6, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_6); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } -else -{ -uint8_t x_64; -lean_dec(x_10); -x_64 = !lean_is_exclusive(x_15); -if (x_64 == 0) +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1() { +_start: { -return x_15; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("object file '", 13, 13); +return x_1; } -else +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2() { +_start: { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_15, 0); -x_66 = lean_ctor_get(x_15, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_15); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' of module ", 12, 12); +return x_1; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" does not exist", 15, 15); +return x_1; } -else +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; -x_68 = lean_ctor_get(x_6, 0); -x_69 = lean_ctor_get(x_6, 1); -x_70 = lean_ctor_get(x_6, 2); -x_71 = lean_ctor_get(x_6, 3); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_6); -x_72 = lean_ctor_get(x_1, 0); -lean_inc(x_72); +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_st_ref_take(x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_6, 0); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); lean_dec(x_1); -lean_inc(x_72); -x_73 = l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_68, x_72); -x_74 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_69); -lean_ctor_set(x_74, 2, x_70); -lean_ctor_set(x_74, 3, x_71); -x_75 = lean_st_ref_set(x_3, x_74, x_7); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_77 = 1; -lean_inc(x_72); -x_78 = l_Lean_findOLean(x_72, x_77, x_76); -if (lean_obj_tag(x_78) == 0) +lean_inc(x_10); +x_11 = l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_9, x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = lean_st_ref_set(x_3, x_6, x_7); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_10); +x_14 = l_Lean_findOLean(x_10, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_read_module_data(x_79, x_80); -lean_dec(x_79); -if (lean_obj_tag(x_81) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_System_FilePath_pathExists(x_15, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_82, 1); -lean_inc(x_85); -lean_dec(x_82); -x_86 = lean_ctor_get(x_84, 0); -lean_inc(x_86); -x_87 = l_Lean_importModulesCore(x_86, x_3, x_83); -lean_dec(x_86); -if (lean_obj_tag(x_87) == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = lean_st_ref_take(x_3, x_88); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -lean_dec(x_89); -x_92 = lean_ctor_get(x_90, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_90, 1); -lean_inc(x_93); -x_94 = lean_ctor_get(x_90, 2); -lean_inc(x_94); -x_95 = lean_ctor_get(x_90, 3); -lean_inc(x_95); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - lean_ctor_release(x_90, 2); - lean_ctor_release(x_90, 3); - x_96 = x_90; -} else { - lean_dec_ref(x_90); - x_96 = lean_box(0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; +x_23 = lean_string_append(x_22, x_15); +lean_dec(x_15); +x_24 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; +x_25 = lean_string_append(x_23, x_24); +x_26 = 1; +x_27 = l_Lean_Name_toString(x_10, x_26); +x_28 = lean_string_append(x_25, x_27); +lean_dec(x_27); +x_29 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_tag(x_17, 1); +lean_ctor_set(x_17, 0, x_31); +return x_17; } -x_97 = lean_array_push(x_93, x_72); -x_98 = lean_array_push(x_94, x_84); -x_99 = lean_array_push(x_95, x_85); -if (lean_is_scalar(x_96)) { - x_100 = lean_alloc_ctor(0, 4, 0); -} else { - x_100 = x_96; +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_dec(x_17); +x_33 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; +x_34 = lean_string_append(x_33, x_15); +lean_dec(x_15); +x_35 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; +x_36 = lean_string_append(x_34, x_35); +x_37 = 1; +x_38 = l_Lean_Name_toString(x_10, x_37); +x_39 = lean_string_append(x_36, x_38); +lean_dec(x_38); +x_40 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; +x_41 = lean_string_append(x_39, x_40); +x_42 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_32); +return x_43; } -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_97); -lean_ctor_set(x_100, 2, x_98); -lean_ctor_set(x_100, 3, x_99); -x_101 = lean_st_ref_set(x_3, x_100, x_91); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); } -x_104 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_17, 1); +lean_inc(x_44); +lean_dec(x_17); +x_45 = lean_box(0); +x_46 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_15, x_10, x_45, x_3, x_44); +lean_dec(x_15); +return x_46; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_85); -lean_dec(x_84); -lean_dec(x_72); -x_106 = lean_ctor_get(x_87, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_87, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_108 = x_87; -} else { - lean_dec_ref(x_87); - x_108 = lean_box(0); +uint8_t x_47; +lean_dec(x_10); +x_47 = !lean_is_exclusive(x_14); +if (x_47 == 0) +{ +return x_14; } -if (lean_is_scalar(x_108)) { - x_109 = lean_alloc_ctor(1, 2, 0); -} else { - x_109 = x_108; +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_14, 0); +x_49 = lean_ctor_get(x_14, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_14); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } -lean_ctor_set(x_109, 0, x_106); -lean_ctor_set(x_109, 1, x_107); -return x_109; } } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -lean_dec(x_72); -x_110 = lean_ctor_get(x_81, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_81, 1); -lean_inc(x_111); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_112 = x_81; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_51 = lean_ctor_get(x_6, 0); +x_52 = lean_ctor_get(x_6, 1); +x_53 = lean_ctor_get(x_6, 2); +x_54 = lean_ctor_get(x_6, 3); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_6); +x_55 = lean_ctor_get(x_1, 0); +lean_inc(x_55); +lean_dec(x_1); +lean_inc(x_55); +x_56 = l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(x_51, x_55); +x_57 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_52); +lean_ctor_set(x_57, 2, x_53); +lean_ctor_set(x_57, 3, x_54); +x_58 = lean_st_ref_set(x_3, x_57, x_7); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +lean_inc(x_55); +x_60 = l_Lean_findOLean(x_55, x_59); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_System_FilePath_pathExists(x_61, x_62); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_unbox(x_64); +lean_dec(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_67 = x_63; } else { - lean_dec_ref(x_81); - x_112 = lean_box(0); + lean_dec_ref(x_63); + x_67 = lean_box(0); } -if (lean_is_scalar(x_112)) { - x_113 = lean_alloc_ctor(1, 2, 0); +x_68 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; +x_69 = lean_string_append(x_68, x_61); +lean_dec(x_61); +x_70 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; +x_71 = lean_string_append(x_69, x_70); +x_72 = 1; +x_73 = l_Lean_Name_toString(x_55, x_72); +x_74 = lean_string_append(x_71, x_73); +lean_dec(x_73); +x_75 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; +x_76 = lean_string_append(x_74, x_75); +x_77 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_77, 0, x_76); +if (lean_is_scalar(x_67)) { + x_78 = lean_alloc_ctor(1, 2, 0); } else { - x_113 = x_112; + x_78 = x_67; + lean_ctor_set_tag(x_78, 1); } -lean_ctor_set(x_113, 0, x_110); -lean_ctor_set(x_113, 1, x_111); -return x_113; +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_66); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_63, 1); +lean_inc(x_79); +lean_dec(x_63); +x_80 = lean_box(0); +x_81 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_61, x_55, x_80, x_3, x_79); +lean_dec(x_61); +return x_81; } } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_72); -x_114 = lean_ctor_get(x_78, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_78, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_116 = x_78; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_55); +x_82 = lean_ctor_get(x_60, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_60, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_84 = x_60; } else { - lean_dec_ref(x_78); - x_116 = lean_box(0); + lean_dec_ref(x_60); + x_84 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); } else { - x_117 = x_116; + x_85 = x_84; } -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; } } } @@ -11970,7 +11987,7 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_9, x_17, x_5, x_13); +x_18 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(x_9, x_17, x_5, x_13); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; @@ -12133,11 +12150,22 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; @@ -14459,7 +14487,7 @@ x_7 = l_Lean_withImportModules___rarg(x_1, x_2, x_6, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -14482,7 +14510,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -14520,7 +14548,7 @@ size_t x_15; size_t x_16; lean_object* x_17; x_15 = 0; x_16 = lean_usize_of_nat(x_7); lean_dec(x_7); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2(x_6, x_15, x_16, x_4); +x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2(x_6, x_15, x_16, x_4); lean_dec(x_6); x_2 = x_11; x_4 = x_17; @@ -14534,7 +14562,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -14561,13 +14589,13 @@ size_t x_7; size_t x_8; lean_object* x_9; x_7 = 0; x_8 = lean_usize_of_nat(x_3); lean_dec(x_3); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3(x_2, x_7, x_8, x_1); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3(x_2, x_7, x_8, x_1); return x_9; } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(lean_object* x_1) { _start: { uint8_t x_2; @@ -14605,7 +14633,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -14631,7 +14659,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -14640,7 +14668,7 @@ x_4 = l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -14649,7 +14677,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -14660,15 +14688,15 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_dec(x_2); -x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1; -x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(x_5, x_1); +x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1; +x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(x_5, x_1); x_7 = 1; x_8 = l_Lean_mkEmptyEnvironment___lambda__1___closed__2; x_9 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_9, 0, x_6); lean_ctor_set(x_9, 1, x_8); lean_ctor_set_uint8(x_9, sizeof(void*)*2, x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(x_9); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(x_9); return x_10; } else @@ -14679,15 +14707,15 @@ if (x_11 == 0) { lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_2); -x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1; -x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(x_12, x_1); +x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1; +x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(x_12, x_1); x_14 = 1; x_15 = l_Lean_mkEmptyEnvironment___lambda__1___closed__2; x_16 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_16, 0, x_13); lean_ctor_set(x_16, 1, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_14); -x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(x_16); +x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(x_16); return x_17; } else @@ -14696,23 +14724,23 @@ size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_2 x_18 = 0; x_19 = lean_usize_of_nat(x_2); lean_dec(x_2); -x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5(x_1, x_18, x_19, x_3); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5(x_1, x_18, x_19, x_3); x_21 = l_Lean_mkHashMapImp___rarg(x_20); lean_dec(x_20); -x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(x_21, x_1); +x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(x_21, x_1); x_23 = 1; x_24 = l_Lean_mkEmptyEnvironment___lambda__1___closed__2; x_25 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_25, 0, x_22); lean_ctor_set(x_25, 1, x_24); lean_ctor_set_uint8(x_25, sizeof(void*)*2, x_23); -x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__4(x_25); +x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__4(x_25); return x_26; } } } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1() { _start: { lean_object* x_1; @@ -14720,17 +14748,17 @@ x_1 = lean_mk_string_unchecked("namespacesExt", 13, 13); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___auto____x40_Lean_Environment___hyg_2607____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3() { _start: { lean_object* x_1; @@ -14739,30 +14767,30 @@ lean_closure_set(x_1, 0, lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -14771,16 +14799,16 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6; x_3 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -14788,12 +14816,12 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__2(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__2(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -14801,21 +14829,21 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__3(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__3(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__1(x_1, x_2); +x_3 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -14823,16 +14851,16 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6347____spec__5(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6443____spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2(x_1); +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2(x_1); lean_dec(x_1); return x_2; } @@ -14873,7 +14901,7 @@ x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__1; return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -14883,11 +14911,11 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6476_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6572_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1; x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); return x_3; } @@ -17945,6 +17973,12 @@ l_Lean_throwAlreadyImported___rarg___closed__7 = _init_l_Lean_throwAlreadyImport lean_mark_persistent(l_Lean_throwAlreadyImported___rarg___closed__7); l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1); +l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2); +l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3); l_Lean_finalizeImport___lambda__2___closed__1 = _init_l_Lean_finalizeImport___lambda__2___closed__1(); lean_mark_persistent(l_Lean_finalizeImport___lambda__2___closed__1); l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___closed__1(); @@ -17957,21 +17991,21 @@ l_Lean_importModules___closed__1 = _init_l_Lean_importModules___closed__1(); lean_mark_persistent(l_Lean_importModules___closed__1); l_Lean_importModules___boxed__const__1 = _init_l_Lean_importModules___boxed__const__1(); lean_mark_persistent(l_Lean_importModules___boxed__const__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6347____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6347_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6443____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6443_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_namespacesExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_namespacesExt); @@ -17983,9 +18017,9 @@ l_Lean_Kernel_instInhabitedDiagnostics___closed__1 = _init_l_Lean_Kernel_instInh lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics___closed__1); l_Lean_Kernel_instInhabitedDiagnostics = _init_l_Lean_Kernel_instInhabitedDiagnostics(); lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics); -l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6476____closed__1); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6476_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6572____closed__1); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6572_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_diagExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_diagExt); diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index 070cef49d331..b561a426f997 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -8313,14 +8313,16 @@ return x_2; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 2); x_7 = lean_ctor_get(x_2, 3); x_8 = lean_ctor_get(x_2, 4); x_9 = lean_ctor_get(x_2, 5); x_10 = lean_ctor_get(x_2, 6); -x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_11 = lean_ctor_get(x_2, 7); +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*8); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -8328,16 +8330,17 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_1); -lean_ctor_set(x_12, 2, x_6); -lean_ctor_set(x_12, 3, x_7); -lean_ctor_set(x_12, 4, x_8); -lean_ctor_set(x_12, 5, x_9); -lean_ctor_set(x_12, 6, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*7, x_11); -return x_12; +x_13 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_1); +lean_ctor_set(x_13, 2, x_6); +lean_ctor_set(x_13, 3, x_7); +lean_ctor_set(x_13, 4, x_8); +lean_ctor_set(x_13, 5, x_9); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*8, x_12); +return x_13; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c index 5de66a8271d0..93587bdc6181 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c +++ b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c @@ -13,31 +13,33 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_get_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Meta_Linear_Cnstr_getBound___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstrKind___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2702____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3; static lean_object* l_Lean_Meta_Linear_instReprVar___closed__1; uint8_t l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_assignment___default; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstr; static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_resolve(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Int_instInhabited; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprAssumptionId; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Linear_getBestBound_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqJustification(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8; static lean_object* l_Lean_Meta_Linear_instInhabitedJustification___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_currVar___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstrKind; @@ -49,24 +51,25 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_currVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size___boxed(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_Cnstr_isStrict(lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqPoly___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6; uint8_t lean_usize_dec_eq(size_t, size_t); extern uint8_t l_instInhabitedBool; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209____boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instInhabitedCnstrKind; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_420____closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_CnstrKind_ofNat(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____boxed(lean_object*, lean_object*); @@ -78,75 +81,77 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Linear_Cnstr_isUnsat___spec__1(l static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstrKind(uint8_t, uint8_t); static lean_object* l_Lean_Meta_Linear_instOrdVar___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19; static lean_object* l_Lean_Meta_Linear_Assignment_val___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstr; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_getBestBound_x3f(lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedJustification; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7; lean_object* l_Lean_Rat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2; static lean_object* l_Lean_Meta_Linear_instInhabitedState___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6; static lean_object* l_Lean_Meta_Linear_instBEqCnstrKind___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__19; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_AssumptionId_id___default; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_420_(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2794____boxed(lean_object*, lean_object*); static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__3; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4; static lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Linear_getBestBound_x3f___spec__1(lean_object*, uint8_t, uint8_t, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_420____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2794_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Rat_inv(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableLtVar(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594_(uint8_t, uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instOrdVar; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1; static lean_object* l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__6; lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f(lean_object*, uint8_t, lean_object*, uint8_t); static lean_object* l_Lean_Meta_Linear_instReprCnstrKind___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__12; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810____boxed(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprAssumptionId___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstr___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestUpperBound_x3f___boxed(lean_object*); static lean_object* l_Lean_Meta_Linear_instReprJustification___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_scale(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612_(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_solve(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Meta_Linear_Cnstr_getBound___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__2; @@ -155,8 +160,7 @@ lean_object* l_Int_repr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprJustification; LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssumptionId; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqJustification___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedVar; @@ -170,40 +174,35 @@ static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__4; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqPoly(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestUpperBound_x3f(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssignment; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size(lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprPoly; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2612_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___boxed(lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_420____closed__2; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstrKind; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_shrink___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25_(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16; LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_scale___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__4; @@ -216,27 +215,23 @@ lean_object* lean_string_length(lean_object*); lean_object* l_Int_instDecidableEq___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqAssumptionId(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__10; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instBEqCnstr___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestLowerBound_x3f(lean_object*); static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__9; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3; lean_object* l_instDecidableEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instLTVar; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestLowerBound_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isStrict___boxed(lean_object*); lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2702_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprCnstr___closed__1; uint8_t l_Lean_Rat_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedCnstr; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); @@ -250,21 +245,24 @@ uint8_t lean_int_dec_lt(lean_object*, lean_object*); uint8_t l_Array_instDecidableEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqAssumptionId___boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810_(lean_object*, lean_object*); static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Linear_Poly_scale___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_420____closed__3; static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_resolve___closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__18; lean_object* l_Lean_Rat_floor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add(lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_Lean_Rat_ceil(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_resolve___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedPoly; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__3; @@ -272,20 +270,20 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__9; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_get_x3f___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__5; lean_object* lean_int_add(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____boxed(lean_object*, lean_object*); uint8_t l_Lean_Rat_isInt(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Rat_sub(lean_object*, lean_object*); @@ -295,12 +293,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableLtVar___boxed(lean_obje LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_ofNat___boxed(lean_object*); lean_object* lean_int_neg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_val___default; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprVar; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__6; @@ -314,14 +311,17 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqVar___boxed(lean_obje static lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVar___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_toCtorIdx(uint8_t); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__13; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____closed__17; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8; static lean_object* _init_l_Lean_Meta_Linear_instInhabitedVar() { _start: { @@ -2914,7 +2914,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2940,7 +2940,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_4, x_8); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209_(x_4, x_8); if (x_13 == 0) { uint8_t x_14; @@ -2992,11 +2992,11 @@ return x_22; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3007,7 +3007,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqJustification___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209____boxed), 2, 0); return x_1; } } @@ -3019,7 +3019,7 @@ x_1 = l_Lean_Meta_Linear_instBEqJustification___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1() { _start: { lean_object* x_1; @@ -3027,21 +3027,21 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.Justification.combine", 38, 38) return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3049,7 +3049,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -3058,7 +3058,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -3067,7 +3067,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6() { _start: { lean_object* x_1; @@ -3075,21 +3075,21 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.Justification.assumption", 41, return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3097,7 +3097,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3116,20 +3116,20 @@ x_7 = lean_unsigned_to_nat(1024u); x_8 = lean_nat_dec_le(x_7, x_2); x_9 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_382____spec__2___closed__1; x_10 = lean_int_dec_lt(x_3, x_9); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(x_4, x_7); +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(x_4, x_7); x_12 = lean_int_dec_lt(x_5, x_9); -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(x_6, x_7); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(x_6, x_7); if (x_8 == 0) { lean_object* x_48; -x_48 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; +x_48 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; x_14 = x_48; goto block_47; } else { lean_object* x_49; -x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; +x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; x_14 = x_49; goto block_47; } @@ -3160,7 +3160,7 @@ goto block_41; block_41: { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3; +x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -3239,14 +3239,14 @@ lean_dec(x_1); x_51 = lean_unsigned_to_nat(1024u); x_52 = lean_nat_dec_le(x_51, x_2); x_53 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2702_(x_50, x_51); -x_54 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8; +x_54 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8; x_55 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); if (x_52 == 0) { lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; -x_56 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; +x_56 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; x_57 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_57, 0, x_56); lean_ctor_set(x_57, 1, x_55); @@ -3260,7 +3260,7 @@ return x_60; else { lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; -x_61 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; +x_61 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; x_62 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_62, 1, x_55); @@ -3274,11 +3274,11 @@ return x_65; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3287,7 +3287,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprJustification___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____boxed), 2, 0); return x_1; } } @@ -3478,7 +3478,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -3490,7 +3490,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -3498,7 +3498,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594_(x_3, x_4); +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -3507,7 +3507,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595____boxed), 2, 0); return x_1; } } @@ -3519,7 +3519,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstrKind___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1() { _start: { lean_object* x_1; @@ -3527,33 +3527,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.eq", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3561,23 +3561,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3585,7 +3585,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7() { _start: { lean_object* x_1; @@ -3593,33 +3593,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.div", 30, 30); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3627,23 +3627,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3651,7 +3651,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13() { _start: { lean_object* x_1; @@ -3659,33 +3659,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.lt", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3693,23 +3693,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3717,7 +3717,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19() { _start: { lean_object* x_1; @@ -3725,33 +3725,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.le", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3759,23 +3759,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3783,7 +3783,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -3795,14 +3795,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6; +x_7 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -3815,14 +3815,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -3835,14 +3835,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18; +x_19 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -3855,14 +3855,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22; +x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24; +x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -3870,13 +3870,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612_(x_3, x_2); +x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -3885,7 +3885,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____boxed), 2, 0); return x_1; } } @@ -3921,7 +3921,7 @@ x_1 = l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -3970,11 +3970,11 @@ return x_17; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3985,7 +3985,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstr(lean_object* x_1, le _start: { uint8_t x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3809_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3810_(x_1, x_2); return x_3; } } @@ -4000,7 +4000,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -4012,7 +4012,7 @@ x_7 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); x_10 = lean_ctor_get(x_2, 2); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3594_(x_3, x_7); +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3595_(x_3, x_7); if (x_11 == 0) { uint8_t x_12; @@ -4042,18 +4042,18 @@ return x_16; else { uint8_t x_17; -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_6, x_10); +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3209_(x_6, x_10); return x_17; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -4064,7 +4064,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4070____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4071____boxed), 2, 0); return x_1; } } @@ -4076,7 +4076,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstr___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1() { _start: { lean_object* x_1; @@ -4084,33 +4084,33 @@ x_1 = lean_mk_string_unchecked("kind", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3; x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4118,7 +4118,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -4127,7 +4127,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6() { _start: { lean_object* x_1; @@ -4135,17 +4135,17 @@ x_1 = lean_mk_string_unchecked("lhs", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8() { _start: { lean_object* x_1; @@ -4153,17 +4153,17 @@ x_1 = lean_mk_string_unchecked("rhs", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10() { _start: { lean_object* x_1; @@ -4171,24 +4171,24 @@ x_1 = lean_mk_string_unchecked("jst", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613_(x_3, x_4); +x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -4196,7 +4196,7 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4; +x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -4208,7 +4208,7 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7; +x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -4235,7 +4235,7 @@ lean_ctor_set(x_26, 1, x_12); x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_14); -x_28 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9; +x_28 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -4249,7 +4249,7 @@ x_33 = lean_int_dec_lt(x_31, x_32); x_34 = lean_ctor_get(x_1, 2); lean_inc(x_34); lean_dec(x_1); -x_35 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337_(x_34, x_4); +x_35 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338_(x_34, x_4); x_36 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_36, 0, x_22); lean_ctor_set(x_36, 1, x_35); @@ -4278,7 +4278,7 @@ lean_ctor_set(x_43, 1, x_12); x_44 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_44, 0, x_43); lean_ctor_set(x_44, 1, x_14); -x_45 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11; +x_45 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11; x_46 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); @@ -4328,7 +4328,7 @@ lean_ctor_set(x_62, 1, x_12); x_63 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_63, 0, x_62); lean_ctor_set(x_63, 1, x_14); -x_64 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11; +x_64 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11; x_65 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_65, 0, x_63); lean_ctor_set(x_65, 1, x_64); @@ -4357,11 +4357,11 @@ return x_74; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4370,7 +4370,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____boxed), 2, 0); return x_1; } } @@ -5320,7 +5320,7 @@ static lean_object* _init_l_Lean_Meta_Linear_pickAssignment_x3f___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5341,7 +5341,7 @@ static lean_object* _init_l_Lean_Meta_Linear_pickAssignment_x3f___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6654,22 +6654,22 @@ l_Lean_Meta_Linear_instBEqJustification___closed__1 = _init_l_Lean_Meta_Linear_i lean_mark_persistent(l_Lean_Meta_Linear_instBEqJustification___closed__1); l_Lean_Meta_Linear_instBEqJustification = _init_l_Lean_Meta_Linear_instBEqJustification(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqJustification); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3337____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3338____closed__8); l_Lean_Meta_Linear_instReprJustification___closed__1 = _init_l_Lean_Meta_Linear_instReprJustification___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprJustification___closed__1); l_Lean_Meta_Linear_instReprJustification = _init_l_Lean_Meta_Linear_instReprJustification(); @@ -6681,54 +6681,54 @@ l_Lean_Meta_Linear_instBEqCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instB lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind___closed__1); l_Lean_Meta_Linear_instBEqCnstrKind = _init_l_Lean_Meta_Linear_instBEqCnstrKind(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__11); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__12); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__13); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__14); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__15); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__16); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__17); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__18); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__19); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__20); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__21); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__22); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__23); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3612____closed__24); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__12); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__13); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__14); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__15); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__16); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__17); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__18); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__19); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__20); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__21); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__22); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__23); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3613____closed__24); l_Lean_Meta_Linear_instReprCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstrKind___closed__1); l_Lean_Meta_Linear_instReprCnstrKind = _init_l_Lean_Meta_Linear_instReprCnstrKind(); @@ -6741,28 +6741,28 @@ l_Lean_Meta_Linear_instBEqCnstr___closed__1 = _init_l_Lean_Meta_Linear_instBEqCn lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr___closed__1); l_Lean_Meta_Linear_instBEqCnstr = _init_l_Lean_Meta_Linear_instBEqCnstr(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4176____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4177____closed__11); l_Lean_Meta_Linear_instReprCnstr___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstr___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstr___closed__1); l_Lean_Meta_Linear_instReprCnstr = _init_l_Lean_Meta_Linear_instReprCnstr(); diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index aff5efaccaa2..3c14d20d8524 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -30,6 +30,7 @@ lean_object* l_Lean_Parser_many1_formatter(lean_object*, lean_object*, lean_obje static lean_object* l_Lean_Parser_Command_register__tactic__tag_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39; static lean_object* l_Lean_Parser_Command_openRenaming___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__7; @@ -43,6 +44,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesi static lean_object* l_Lean_Parser_Command_nonrec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_formatter__1(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44; static lean_object* l_Lean_Parser_Command_tactic__extension___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange__1___closed__1; @@ -60,6 +62,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_formatter(lean_object*, le static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__6; static lean_object* l_Lean_Parser_Command_openSimple___closed__6; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43; static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__8; @@ -95,7 +98,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor_parenthesizer(lean_obj LEAN_EXPORT lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim; static lean_object* l_Lean_Parser_Command_ctor___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54; static lean_object* l_Lean_Parser_Command_namespace___closed__3; static lean_object* l_Lean_Parser_Command_protected___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__6; @@ -219,8 +221,8 @@ static lean_object* l_Lean_Parser_Command_printEqns_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenamingItem_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange__1___closed__6; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__16; @@ -247,7 +249,6 @@ static lean_object* l_Lean_Parser_Command_declModifiers___closed__14; static lean_object* l_Lean_Parser_Tactic_open___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__1; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -292,7 +293,6 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__10; static lean_object* l_Lean_Parser_Command_eval_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__14; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__15; @@ -325,12 +325,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange__1___close static lean_object* l_Lean_Parser_Command_addDocString_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_docString__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange__1___closed__5; static lean_object* l_Lean_Parser_Command_structCtor___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_formatter__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_nonrec___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35; static lean_object* l_Lean_Parser_Term_open___closed__4; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; @@ -350,6 +350,7 @@ static lean_object* l_Lean_Parser_Command_openSimple_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_opaque_parenthesizer__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_openScoped___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52; static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer__1___closed__1; @@ -361,7 +362,6 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___c static lean_object* l_Lean_Parser_Command_declSig___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__6; static lean_object* l_Lean_Parser_Command_deriving___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__3; static lean_object* l_Lean_Parser_Command_universe___closed__6; @@ -370,6 +370,7 @@ static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__3; static lean_object* l_Lean_Parser_Command_include___closed__5; @@ -410,7 +411,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__6; extern lean_object* l_Lean_Parser_Term_structInst; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__3; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29; static lean_object* l_Lean_Parser_Command_include___closed__4; static lean_object* l_Lean_Parser_Command_partial___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -442,7 +442,6 @@ lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Parser_Command_definition_formatter___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_check_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_include___closed__10; static lean_object* l_Lean_Parser_Command_openOnly___closed__5; static lean_object* l_Lean_Parser_Command_attribute___closed__15; static lean_object* l_Lean_Parser_Command_section___closed__3; @@ -455,7 +454,6 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__10; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__4; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41; static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange__1___closed__2; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__3; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__5; @@ -463,6 +461,7 @@ extern lean_object* l_Lean_Parser_Term_binderDefault; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_import_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53; static lean_object* l_Lean_Parser_Command_attribute___closed__16; lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange__1(lean_object*); @@ -480,12 +479,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quot__1___closed__1; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor; static lean_object* l_Lean_Parser_Command_init__quot___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check_parenthesizer__1(lean_object*); lean_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38; static lean_object* l_Lean_Parser_Command_structureTk_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declBody___closed__2; @@ -499,7 +496,6 @@ static lean_object* l_Lean_Parser_Command_definition___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_structure___closed__2; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2; static lean_object* l_Lean_Parser_Command_check__failure___closed__1; static lean_object* l_Lean_Parser_Command_init__quot_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId___lambda__1___closed__4___boxed__const__1; @@ -508,6 +504,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange__1___closed__3; static lean_object* l_Lean_Parser_Command_check___closed__2; lean_object* l_Lean_Parser_recover(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51; static lean_object* l_Lean_Parser_Command_tactic__extension_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_open; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange__1___closed__3; @@ -549,6 +546,7 @@ static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__9; static lean_object* l_Lean_Parser_Command_deriving___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_Command_declValEqns_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_print___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -691,6 +689,7 @@ static lean_object* l_Lean_Parser_Command_declValSimple___closed__4; static lean_object* l_Lean_Parser_Command_example___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__4; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__3; static lean_object* l_Lean_Parser_Command_import_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__2; @@ -710,7 +709,9 @@ static lean_object* l_Lean_Parser_Command_computedField___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_protected_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_variable___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_formatter__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9; lean_object* l_Lean_Parser_incQuotDepth_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__8; extern lean_object* l_Lean_Parser_Term_matchAltsWhereDecls; @@ -748,6 +749,7 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__14; static lean_object* l_Lean_Parser_Command_instance___closed__6; static lean_object* l_Lean_Parser_Command_example_parenthesizer___closed__2; lean_object* l_Lean_Parser_orelse(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__19; static lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_section_parenthesizer___closed__2; @@ -767,17 +769,15 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDefDeriving; static lean_object* l_Lean_Parser_Command_optionValue___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange__1___closed__1; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputableSection; static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWs___closed__1; static lean_object* l_Lean_Parser_Term_quot___closed__15; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__9; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange__1___closed__1; static lean_object* l_Lean_Parser_Command_declId___closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15; static lean_object* l_Lean_Parser_Command_optNamedPrio___closed__1; static lean_object* l_Lean_Parser_Command_printAxioms___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter__1(lean_object*); @@ -791,7 +791,6 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__4; static lean_object* l_Lean_Parser_Command_optDeriving___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_extends_parenthesizer__1(lean_object*); @@ -812,8 +811,8 @@ static lean_object* l_Lean_Parser_Command_moduleDoc___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_theorem; static lean_object* l_Lean_Parser_Command_namedPrio___closed__17; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__11; static lean_object* l_Lean_Parser_Command_set__option___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_formatter__1___closed__2; @@ -884,7 +883,6 @@ static lean_object* l_Lean_Parser_Command_partial_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_parenthesizer__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structFields___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10; lean_object* l_Lean_Parser_Term_attrInstance_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_end_docString__1___closed__1; @@ -906,19 +904,19 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValSimple_formatter(lean_obje static lean_object* l_Lean_Parser_Term_open___closed__3; static lean_object* l_Lean_Parser_Command_declVal_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5; static lean_object* l_Lean_Parser_Command_printTacTags_formatter___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initialize_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_exit___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_include__1(lean_object*); static lean_object* l_Lean_Parser_Command_check__failure___closed__5; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_eoi___closed__5; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_openOnly_formatter__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField; static lean_object* l_Lean_Parser_Command_attribute___closed__17; @@ -1063,6 +1061,7 @@ static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__15; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_printAxioms___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3; static lean_object* l_Lean_Parser_Command_tactic__extension___closed__8; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_open_formatter___closed__1; @@ -1219,6 +1218,7 @@ lean_object* l_Lean_Parser_Term_binderIdent_formatter(lean_object*, lean_object* static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_register__tactic__tag_declRange__1___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__15; static lean_object* l_Lean_Parser_Command_optionValue_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__11; @@ -1236,7 +1236,6 @@ static lean_object* l_Lean_Parser_Command_check__failure___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange__1___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_set__option___closed__8; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__1; @@ -1252,11 +1251,11 @@ static lean_object* l_Lean_Parser_Command_extends___closed__1; static lean_object* l_Lean_Parser_Command_declId___closed__19; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__9; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30; static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange__1___closed__2; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3041_(lean_object*); static lean_object* l_Lean_Parser_Command_evalBang___closed__9; static lean_object* l_Lean_Parser_Command_classTk___closed__6; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__1; @@ -1277,6 +1276,7 @@ static lean_object* l_Lean_Parser_Term_open___closed__1; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structure___closed__20; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange__1(lean_object*); @@ -1392,7 +1392,6 @@ static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___close static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24; static lean_object* l_Lean_Parser_Command_declId___closed__3; static lean_object* l_Lean_Parser_Command_structureTk_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_open_formatter__1___closed__2; @@ -1438,7 +1437,6 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_deriving___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__9; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3048_(lean_object*); static lean_object* l_Lean_Parser_Command_open_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__3; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__3; @@ -1463,6 +1461,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__18; static lean_object* l_Lean_Parser_Command_printAxioms_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_namespace___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange__1___closed__1; static lean_object* l_Lean_Parser_Command_private_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__5; @@ -1480,7 +1479,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_declRange__1 static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__7; static lean_object* l_Lean_Parser_Tactic_set__option___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_formatter__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_unsafe_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__1; @@ -1563,6 +1561,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter_ static lean_object* l_Lean_Parser_Command_eraseAttr___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eoi_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_derivingClasses___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId___closed__17; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange__1___closed__3; @@ -1572,6 +1571,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer__1___c LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedField; static lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45; static lean_object* l_Lean_Parser_Command_namedPrio___closed__3; static lean_object* l_Lean_Parser_Command_export___closed__12; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__9; @@ -1599,6 +1599,7 @@ static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__6; static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openHiding_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_namedPrio___closed__15; @@ -1655,7 +1656,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_import_formatter(lean_object*, le lean_object* l_Lean_ppLine_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__13; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073_(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_unsafe_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__7; extern lean_object* l_Lean_Parser_Term_letDecl; @@ -1726,6 +1726,7 @@ static lean_object* l_Lean_Parser_Command_namedPrio___closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_declVal___closed__5; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27; lean_object* l_Lean_Parser_symbol(lean_object*); static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_namedPrio___closed__11; @@ -1733,7 +1734,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_parenthesizer__1_ static lean_object* l___regBuiltin_Lean_Parser_Command_opaque_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__4; lean_object* l_Lean_Parser_Term_typeSpec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3; static lean_object* l_Lean_Parser_Term_precheckedQuot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initializeKeyword_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_definition_formatter___closed__4; @@ -1747,6 +1747,7 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__11; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_instance_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20; static lean_object* l_Lean_Parser_Command_declValSimple___closed__5; static lean_object* l_Lean_Parser_Command_axiom___closed__7; static lean_object* l_Lean_Parser_Command_protected___closed__3; @@ -1777,7 +1778,6 @@ static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__5; static lean_object* l_Lean_Parser_Command_classInductive___closed__14; static lean_object* l_Lean_Parser_Command_printEqns___closed__9; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__13; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52; lean_object* l_Lean_Parser_symbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst___closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter__1(lean_object*); @@ -1814,6 +1814,7 @@ static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_export___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_export___closed__1; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__10; @@ -1860,11 +1861,11 @@ static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__6; static lean_object* l_Lean_Parser_Command_example___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34; static lean_object* l_Lean_Parser_Command_evalBang___closed__5; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35; static lean_object* l_Lean_Parser_Command_openRenaming___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor; static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer__1___closed__2; @@ -1874,7 +1875,6 @@ static lean_object* l_Lean_Parser_Command_openRenaming___closed__12; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_tactic__extension_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42; static lean_object* l_Lean_Parser_Command_openScoped_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_evalBang___closed__4; @@ -1895,6 +1895,7 @@ static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_export; lean_object* l_Lean_PrettyPrinter_Formatter_lookahead_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7; lean_object* l_Lean_Parser_withAntiquot(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declId___closed__6; static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__3; @@ -1913,7 +1914,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer static lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_partial___closed__7; static lean_object* l_Lean_Parser_Command_structCtor___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13; static lean_object* l_Lean_Parser_Command_initialize___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2042,6 +2042,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersT; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_formatter__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declId___closed__15; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30; static lean_object* l_Lean_Parser_Command_nonrec_formatter___closed__3; static lean_object* l_Lean_Parser_Command_include___closed__7; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__11; @@ -2067,6 +2068,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_formatter__1(lean static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_extends___closed__2; static lean_object* l_Lean_Parser_Command_initialize___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38; static lean_object* l_Lean_Parser_Command_import___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_end; static lean_object* l_Lean_Parser_Command_printAxioms___closed__8; @@ -2075,6 +2077,7 @@ static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__8; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__13; static lean_object* l_Lean_Parser_Command_attribute___closed__13; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22; static lean_object* l_Lean_Parser_Command_printTacTags_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_tactic__extension_declRange__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_tactic__extension_declRange__1___closed__2; @@ -2142,6 +2145,7 @@ static lean_object* l_Lean_Parser_Command_whereStructField___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_eval_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__18; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange__1___closed__4; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_open_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2155,6 +2159,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange__1( static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__7; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__9; static lean_object* l_Lean_Parser_Command_definition___closed__8; static lean_object* l_Lean_Parser_Command_computedFields___closed__2; @@ -2178,7 +2183,6 @@ static lean_object* l_Lean_Parser_Term_set__option___closed__9; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter__1___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_section_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2215,7 +2219,6 @@ static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__4; lean_object* l_Lean_Parser_symbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_moduleDoc___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__12; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export__1(lean_object*); @@ -2261,7 +2264,6 @@ static lean_object* l_Lean_Parser_Command_declId_formatter___closed__3; static lean_object* l_Lean_Parser_Term_quot___closed__6; static lean_object* l_Lean_Parser_Term_open_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36; static lean_object* l_Lean_Parser_Command_check___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_example_formatter__1___closed__1; @@ -2281,6 +2283,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString__1(lean static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__9; lean_object* l_Lean_Parser_atomic(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_open_docString__1___closed__1; static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__3; lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2293,6 +2296,7 @@ static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__1 static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_evalBang_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_docString__1(lean_object*); @@ -2329,6 +2333,7 @@ static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed_ static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_mutual___closed__11; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange__1___closed__3; @@ -2383,6 +2388,7 @@ lean_object* l_Lean_ppDedent_formatter(lean_object*, lean_object*, lean_object*, static lean_object* l_Lean_Parser_Command_universe___closed__1; static lean_object* l_Lean_Parser_Command_derivingClasses___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28; static lean_object* l_Lean_Parser_Command_declId___closed__8; static lean_object* l_Lean_Parser_Command_openOnly___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange__1___closed__2; @@ -2410,7 +2416,6 @@ static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_end___closed__8; static lean_object* l_Lean_Parser_Command_include_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__2; -static lean_object* l_Lean_Parser_Command_include_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_protected_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__3; @@ -2418,7 +2423,6 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49; static lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange__1___closed__4; static lean_object* l_Lean_Parser_Command_check__failure___closed__4; @@ -2446,6 +2450,7 @@ static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__7; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_exit_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__7; lean_object* l_Lean_ppSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize; @@ -2484,7 +2489,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange__1__ lean_object* l_Lean_Parser_withPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open__1(lean_object*); lean_object* l_Lean_Parser_leadingNode_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__1; static lean_object* l_Lean_Parser_Command_check__failure_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_computedField___closed__10; @@ -2498,7 +2502,6 @@ lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_ static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__6; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23; static lean_object* l_Lean_Parser_Command_abbrev___closed__7; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__8; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -2508,6 +2511,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_printTacTags_formatter(lean_objec static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__5; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14; static lean_object* l_Lean_Parser_Command_inductive___closed__4; static lean_object* l_Lean_Parser_Command_attribute___closed__5; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__1; @@ -2572,6 +2576,7 @@ static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_eval_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_eraseAttr_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__2; @@ -2621,6 +2626,7 @@ static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__1(lean_object*); static lean_object* l_Lean_Parser_Command_instance_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46; static lean_object* l_Lean_Parser_Command_openDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange__1___closed__7; @@ -2710,8 +2716,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_ extern lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__7; lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26; static lean_object* l_Lean_Parser_Command_evalBang_formatter___closed__4; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__6; @@ -2743,7 +2751,6 @@ static lean_object* l_Lean_Parser_Command_addDocString_formatter___closed__5; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__2; static lean_object* l_Lean_Parser_Command_variable___closed__6; static lean_object* l_Lean_Parser_Command_initialize___closed__13; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12; static lean_object* l_Lean_Parser_Command_declModifiers___closed__8; static lean_object* l_Lean_Parser_Command_deriving___closed__12; @@ -2756,7 +2763,6 @@ static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_in___closed__2; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__1; static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6; lean_object* l_Lean_Parser_identWithPartialTrailingDot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declValSimple___closed__11; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__4; @@ -2782,7 +2788,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_docStri static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDefDeriving_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__3; @@ -2827,6 +2835,7 @@ static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__4; static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namespace_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__5; static lean_object* l_Lean_Parser_Command_ctor___closed__11; @@ -2923,6 +2932,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange__1___c static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__17; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066_(lean_object*); static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declBody_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_parenthesizer__1___closed__2; @@ -3005,7 +3015,6 @@ static lean_object* l_Lean_Parser_Command_structure___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4; static lean_object* l_Lean_Parser_Command_eval___closed__8; static lean_object* l_Lean_Parser_Command_register__tactic__tag_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange__1___closed__5; @@ -3105,6 +3114,7 @@ static lean_object* l_Lean_Parser_Command_classTk___closed__4; static lean_object* l_Lean_Parser_Command_axiom___closed__5; static lean_object* l_Lean_Parser_Term_open_formatter___closed__5; static lean_object* l_Lean_Parser_Command_declModifiers___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_printTacTags_formatter___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__3; @@ -3117,7 +3127,6 @@ static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__6 static lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_declId___lambda__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_attribute__1(lean_object*); @@ -3197,6 +3206,7 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___c static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_declBody_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openScoped___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__10; static lean_object* l_Lean_Parser_Command_computedField___closed__4; static lean_object* l_Lean_Parser_Command_in___closed__6; @@ -3209,7 +3219,6 @@ static lean_object* l_Lean_Parser_Command_declaration___closed__11; lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l_Lean_Parser_Command_inductive___closed__8; lean_object* l_Lean_Parser_Term_structInst_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_open_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__7; @@ -3249,7 +3258,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthe static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end_docString__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_open___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_protected_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_docString__1(lean_object*); @@ -3258,7 +3266,6 @@ lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, lean_object* LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter__1___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_export_formatter___closed__3; @@ -3267,6 +3274,7 @@ static lean_object* l_Lean_Parser_Command_definition_formatter___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__8; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__2; static lean_object* l_Lean_Parser_Command_unsafe_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8; static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structureTk___closed__1; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__2; @@ -3296,7 +3304,6 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Parser_Command_addDocString___closed__1; lean_object* l_Lean_Parser_strLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_universe___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50; static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__5; @@ -3332,7 +3339,6 @@ static lean_object* l_Lean_Parser_Command_instance___closed__4; static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__12; static lean_object* l_Lean_Parser_Command_declId___closed__18; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_docString__1___closed__1; static lean_object* l_Lean_Parser_Command_addDocString_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__9; @@ -3343,7 +3349,6 @@ static lean_object* l_Lean_Parser_Command_end_formatter___closed__1; extern lean_object* l_Lean_Parser_Term_optType; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structSimpleBinder; static lean_object* l_Lean_Parser_Command_structure___closed__19; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_unsafe___closed__5; static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__7; @@ -3408,7 +3413,6 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__3; static lean_object* l_Lean_Parser_Command_extends_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_declRange__1___closed__7; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26; LEAN_EXPORT lean_object* l_Lean_Parser_Command_example; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__4; static lean_object* l_Lean_Parser_Term_set__option___closed__6; @@ -3464,7 +3468,6 @@ static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__20; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriving; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange__1___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11; static lean_object* l_Lean_Parser_Command_openDecl_formatter___closed__4; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_unsafe_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3474,7 +3477,6 @@ static lean_object* l_Lean_Parser_Command_eraseAttr_parenthesizer___closed__3; lean_object* l_Lean_Parser_sepBy1_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_extends___closed__6; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32; static lean_object* l_Lean_Parser_Command_noncomputable___closed__1; static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__4; @@ -3529,6 +3531,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_docString__ static lean_object* l_Lean_Parser_Term_quot_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_export_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_printEqns_declRange__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_quot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__7; @@ -3549,7 +3552,6 @@ lean_object* l_Lean_Parser_Term_bracketedBinder(uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeriving; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange__1___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__5; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__7; static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__2; @@ -3582,7 +3584,6 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_initialize___closed__8; static lean_object* l_Lean_Parser_Command_declaration___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_include_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__25; @@ -3593,6 +3594,7 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___c static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_set__option___closed__4; static lean_object* l_Lean_Parser_Command_variable___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18; static lean_object* l_Lean_Parser_Command_declValSimple___closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__9; @@ -3611,7 +3613,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_namespace; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_print_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange__1___closed__5; @@ -3688,7 +3689,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_formatter(lean_object*, lean static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18; static lean_object* l_Lean_Parser_Command_include_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_printTacTags___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_parenthesizer__1(lean_object*); @@ -3731,7 +3731,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__10 LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_tactic__extension_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_optionValue_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange__1___closed__3; static lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedFields_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3749,9 +3748,7 @@ static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__2; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_private___closed__6; @@ -3767,6 +3764,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_evalBang__1(lean_obj static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_init__quot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_declRange__1___closed__1; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__8; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__9; @@ -3778,7 +3776,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_openRenamingItem_parenthe LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__3; static lean_object* l_Lean_Parser_Command_import_formatter___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17; static lean_object* l___regBuiltin_Lean_Parser_Command_in_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_openRenaming___closed__5; static lean_object* l_Lean_Parser_Command_namespace_formatter___closed__2; @@ -24645,7 +24642,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_section_docString__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("A `section`/`end` pair delimits the scope of `variable`, `open`, `set_option`, and `local` commands.\nSections can be nested. `section ` provides a label to the section that has to appear with the\nmatching `end`. In either case, the `end` can be omitted, in which case the section is closed at the\nend of the file.\n", 318, 318); +x_1 = lean_mk_string_unchecked("A `section`/`end` pair delimits the scope of `variable`, `include, `open`, `set_option`, and `local`\ncommands. Sections can be nested. `section ` provides a label to the section that has to appear\nwith the matching `end`. In either case, the `end` can be omitted, in which case the section is\nclosed at the end of the file.\n", 328, 328); return x_1; } } @@ -25889,7 +25886,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_variable_docString_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Declares one or more typed variables, or modifies whether already-declared variables are\nimplicit.\n\nIntroduces variables that can be used in definitions within the same `namespace` or `section` block.\nWhen a definition mentions a variable, Lean will add it as an argument of the definition. The\n`variable` command is also able to add typeclass parameters. This is useful in particular when\nwriting many definitions that have parameters in common (see below for an example).\n\nVariable declarations have the same flexibility as regular function paramaters. In particular they\ncan be [explicit, implicit][binder docs], or [instance implicit][tpil classes] (in which case they\ncan be anonymous). This can be changed, for instance one can turn explicit variable `x` into an\nimplicit one with `variable {x}`. Note that currently, you should avoid changing how variables are\nbound and declare new variables at the same time; see [issue 2789] for more on this topic.\n\nSee [*Variables and Sections* from Theorem Proving in Lean][tpil vars] for a more detailed\ndiscussion.\n\n[tpil vars]: https://lean-lang.org/theorem_proving_in_lean4/dependent_type_theory.html#variables-and-sections\n(Variables and Sections on Theorem Proving in Lean)\n[tpil classes]: https://lean-lang.org/theorem_proving_in_lean4/type_classes.html\n(Type classes on Theorem Proving in Lean)\n[binder docs]: https://leanprover-community.github.io/mathlib4_docs/Lean/Expr.html#Lean.BinderInfo\n(Documentation for the BinderInfo type)\n[issue 2789]: https://github.com/leanprover/lean4/issues/2789\n(Issue 2789 on github)\n\n## Examples\n\n```lean\nsection\n variable\n {α : Type u} -- implicit\n (a : α) -- explicit\n [instBEq : BEq α] -- instance implicit, named\n [Hashable α] -- instance implicit, anonymous\n\n def isEqual (b : α) : Bool :=\n a == b\n\n #check isEqual\n -- isEqual.{u} {α : Type u} (a : α) [instBEq : BEq α] (b : α) : Bool\n\n variable\n {a} -- `a` is implicit now\n\n def eqComm {b : α} := a == b ↔ b == a\n\n #check eqComm\n -- eqComm.{u} {α : Type u} {a : α} [instBEq : BEq α] {b : α} : Prop\nend\n```\n\nThe following shows a typical use of `variable` to factor out definition arguments:\n\n```lean\nvariable (Src : Type)\n\nstructure Logger where\n trace : List (Src × String)\n#check Logger\n-- Logger (Src : Type) : Type\n\nnamespace Logger\n -- switch `Src : Type` to be implicit until the `end Logger`\n variable {Src}\n\n def empty : Logger Src where\n trace := []\n #check empty\n -- Logger.empty {Src : Type} : Logger Src\n\n variable (log : Logger Src)\n\n def len :=\n log.trace.length\n #check len\n -- Logger.len {Src : Type} (log : Logger Src) : Nat\n\n variable (src : Src) [BEq Src]\n\n -- at this point all of `log`, `src`, `Src` and the `BEq` instance can all become arguments\n\n def filterSrc :=\n log.trace.filterMap\n fun (src', str') => if src' == src then some str' else none\n #check filterSrc\n -- Logger.filterSrc {Src : Type} (log : Logger Src) (src : Src) [inst✝ : BEq Src] : List String\n\n def lenSrc :=\n log.filterSrc src |>.length\n #check lenSrc\n -- Logger.lenSrc {Src : Type} (log : Logger Src) (src : Src) [inst✝ : BEq Src] : Nat\nend Logger\n```\n\n", 3195, 3174); +x_1 = lean_mk_string_unchecked("Declares one or more typed variables, or modifies whether already-declared variables are\n implicit.\n\nIntroduces variables that can be used in definitions within the same `namespace` or `section` block.\nWhen a definition mentions a variable, Lean will add it as an argument of the definition. This is\nuseful in particular when writing many definitions that have parameters in common (see below for an\nexample).\n\nVariable declarations have the same flexibility as regular function paramaters. In particular they\ncan be [explicit, implicit][binder docs], or [instance implicit][tpil classes] (in which case they\ncan be anonymous). This can be changed, for instance one can turn explicit variable `x` into an\nimplicit one with `variable {x}`. Note that currently, you should avoid changing how variables are\nbound and declare new variables at the same time; see [issue 2789] for more on this topic.\n\nIn *theorem bodies* (i.e. proofs), variables are not included based on usage in order to ensure that\nchanges to the proof cannot change the statement of the overall theorem. Instead, variables are only\navailable to the proof if they have been mentioned in the theorem header or in an `include` command\nor are instance implicit and depend only on such variables.\n\nSee [*Variables and Sections* from Theorem Proving in Lean][tpil vars] for a more detailed\ndiscussion.\n\n[tpil vars]:\nhttps://lean-lang.org/theorem_proving_in_lean4/dependent_type_theory.html#variables-and-sections\n(Variables and Sections on Theorem Proving in Lean) [tpil classes]:\nhttps://lean-lang.org/theorem_proving_in_lean4/type_classes.html (Type classes on Theorem Proving in\nLean) [binder docs]:\nhttps://leanprover-community.github.io/mathlib4_docs/Lean/Expr.html#Lean.BinderInfo (Documentation\nfor the BinderInfo type) [issue 2789]: https://github.com/leanprover/lean4/issues/2789 (Issue 2789\non github)\n\n## Examples\n\n```lean\nsection\n variable\n {α : Type u} -- implicit\n (a : α) -- explicit\n [instBEq : BEq α] -- instance implicit, named\n [Hashable α] -- instance implicit, anonymous\n\n def isEqual (b : α) : Bool :=\n a == b\n\n #check isEqual\n -- isEqual.{u} {α : Type u} (a : α) [instBEq : BEq α] (b : α) : Bool\n\n variable\n {a} -- `a` is implicit now\n\n def eqComm {b : α} := a == b ↔ b == a\n\n #check eqComm\n -- eqComm.{u} {α : Type u} {a : α} [instBEq : BEq α] {b : α} : Prop\nend\n```\n\nThe following shows a typical use of `variable` to factor out definition arguments:\n\n```lean\nvariable (Src : Type)\n\nstructure Logger where\n trace : List (Src × String)\n#check Logger\n-- Logger (Src : Type) : Type\n\nnamespace Logger\n -- switch `Src : Type` to be implicit until the `end Logger`\n variable {Src}\n\n def empty : Logger Src where\n trace := []\n #check empty\n -- Logger.empty {Src : Type} : Logger Src\n\n variable (log : Logger Src)\n\n def len :=\n log.trace.length\n #check len\n -- Logger.len {Src : Type} (log : Logger Src) : Nat\n\n variable (src : Src) [BEq Src]\n\n -- at this point all of `log`, `src`, `Src` and the `BEq` instance can all become arguments\n\n def filterSrc :=\n log.trace.filterMap\n fun (src', str') => if src' == src then some str' else none\n #check filterSrc\n -- Logger.filterSrc {Src : Type} (log : Logger Src) (src : Src) [inst✝ : BEq Src] : List String\n\n def lenSrc :=\n log.filterSrc src |>.length\n #check lenSrc\n -- Logger.lenSrc {Src : Type} (log : Logger Src) (src : Src) [inst✝ : BEq Src] : Nat\nend Logger\n```\n\nThe following example demonstrates availability of variables in proofs:\n```lean\nvariable\n {α : Type} -- available in the proof as indirectly mentioned through `a`\n [ToString α] -- available in the proof as `α` is included\n (a : α) -- available in the proof as mentioned in the header\n {β : Type} -- not available in the proof\n [ToString β] -- not available in the proof\n\ntheorem ex : a = a := rfl\n```\nAfter elaboration of the proof, the following warning will be generated to highlight the unused\nhypothesis:\n```\nincluded section variable '[ToString α]' is not used in 'ex', consider excluding it\n```\nIn such cases, the offending variable declaration should be moved down or into a section so that\nonly theorems that do depend on it follow it until the end of the section.\n", 4294, 4266); return x_1; } } @@ -40579,49 +40576,40 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_include___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_noncomputableSection___closed__8; -x_2 = l_Lean_Parser_many1(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_include___closed__7() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_include___closed__5; -x_2 = l_Lean_Parser_Command_include___closed__6; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_include___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_include___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_include___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_include___closed__7; +x_3 = l_Lean_Parser_Command_include___closed__6; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_include___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_include___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_include___closed__3; -x_2 = l_Lean_Parser_Command_include___closed__8; +x_2 = l_Lean_Parser_Command_include___closed__7; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_include___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_include___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_include___closed__2; -x_2 = l_Lean_Parser_Command_include___closed__9; +x_2 = l_Lean_Parser_Command_include___closed__8; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -40630,7 +40618,7 @@ static lean_object* _init_l_Lean_Parser_Command_include() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_include___closed__10; +x_1 = l_Lean_Parser_Command_include___closed__9; return x_1; } } @@ -40651,7 +40639,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_include_docString__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("To be implemented. ", 19, 19); +x_1 = lean_mk_string_unchecked("`include eeny meeny` instructs Lean to include the section `variable`s `eeny` and `meeny` in all\ndeclarations in the remainder of the current section, differing from the default behavior of\nconditionally including variables based on use in the declaration header. `include` is usually\nfollowed by the `in` combinator to limit the inclusion to the subsequent declaration.\n", 371, 371); return x_1; } } @@ -40696,32 +40684,22 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_include_formatter___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_include_formatter___closed__4() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_include_formatter___closed__2; -x_2 = l_Lean_Parser_Command_include_formatter___closed__3; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_include_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_include_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_include___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_include_formatter___closed__4; +x_3 = l_Lean_Parser_Command_include_formatter___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -40734,7 +40712,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_include_formatter(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_include_formatter___closed__1; -x_7 = l_Lean_Parser_Command_include_formatter___closed__5; +x_7 = l_Lean_Parser_Command_include_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -40803,32 +40781,22 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_include_parenthesizer___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_include_parenthesizer___closed__4() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_include_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_include_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_include_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_include_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_include___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_include_parenthesizer___closed__4; +x_3 = l_Lean_Parser_Command_include_parenthesizer___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -40841,7 +40809,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_include_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_include_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_include_parenthesizer___closed__5; +x_7 = l_Lean_Parser_Command_include_parenthesizer___closed__4; x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -41137,7 +41105,7 @@ x_1 = l_Lean_Parser_Command_eoi___closed__5; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3048_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3041_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -41198,7 +41166,7 @@ x_1 = l_Lean_Parser_Command_ctor___closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41208,7 +41176,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2() { _start: { lean_object* x_1; @@ -41216,19 +41184,19 @@ x_1 = lean_mk_string_unchecked("declModifiersF", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -41238,7 +41206,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -41248,7 +41216,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -41258,12 +41226,12 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -41272,7 +41240,7 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8() { _start: { lean_object* x_1; @@ -41280,17 +41248,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10() { _start: { lean_object* x_1; @@ -41298,7 +41266,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11() { _start: { lean_object* x_1; @@ -41306,17 +41274,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13() { _start: { lean_object* x_1; @@ -41324,7 +41292,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14() { _start: { lean_object* x_1; @@ -41332,17 +41300,17 @@ x_1 = lean_mk_string_unchecked("nestedDeclModifiers", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16() { _start: { lean_object* x_1; @@ -41350,19 +41318,19 @@ x_1 = lean_mk_string_unchecked("declModifiersT", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -41372,7 +41340,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19() { _start: { lean_object* x_1; @@ -41380,17 +41348,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21() { _start: { lean_object* x_1; @@ -41398,17 +41366,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41418,7 +41386,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -41428,7 +41396,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -41438,7 +41406,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -41448,7 +41416,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -41458,7 +41426,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41468,7 +41436,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -41478,7 +41446,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -41488,7 +41456,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -41498,7 +41466,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -41508,7 +41476,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41518,7 +41486,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -41528,7 +41496,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -41538,7 +41506,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -41548,7 +41516,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -41558,7 +41526,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41568,7 +41536,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -41578,7 +41546,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -41588,7 +41556,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -41598,7 +41566,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -41608,7 +41576,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -41618,7 +41586,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -41628,7 +41596,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -41638,7 +41606,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -41648,7 +41616,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -41658,7 +41626,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48() { _start: { lean_object* x_1; @@ -41666,29 +41634,29 @@ x_1 = lean_mk_string_unchecked("docComment", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -41698,17 +41666,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -41718,7 +41686,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -41728,15 +41696,15 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4; -x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5; -x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5; +x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) { @@ -41744,8 +41712,8 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); -x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10; -x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10; +x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -41753,8 +41721,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); -x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13; -x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13; +x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -41762,9 +41730,9 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15; -x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17; -x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17; +x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18; x_20 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_5, x_6, x_16); if (lean_obj_tag(x_20) == 0) { @@ -41772,7 +41740,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20; +x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20; x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); if (lean_obj_tag(x_23) == 0) { @@ -41780,7 +41748,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22; +x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22; x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); if (lean_obj_tag(x_26) == 0) { @@ -41788,10 +41756,10 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); -x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23; x_29 = l_Lean_Parser_Command_declId___closed__2; -x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24; -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25; +x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24; +x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25; x_32 = l_Lean_Parser_registerAlias(x_28, x_29, x_30, x_31, x_6, x_27); if (lean_obj_tag(x_32) == 0) { @@ -41799,7 +41767,7 @@ lean_object* x_33; lean_object* x_34; lean_object* x_35; x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26; +x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26; x_35 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_34, x_33); if (lean_obj_tag(x_35) == 0) { @@ -41807,7 +41775,7 @@ lean_object* x_36; lean_object* x_37; lean_object* x_38; x_36 = lean_ctor_get(x_35, 1); lean_inc(x_36); lean_dec(x_35); -x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27; +x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27; x_38 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_37, x_36); if (lean_obj_tag(x_38) == 0) { @@ -41815,10 +41783,10 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28; x_41 = l_Lean_Parser_Command_declSig___closed__2; -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29; -x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30; +x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29; +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30; x_44 = l_Lean_Parser_registerAlias(x_40, x_41, x_42, x_43, x_6, x_39); if (lean_obj_tag(x_44) == 0) { @@ -41826,7 +41794,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_44, 1); lean_inc(x_45); lean_dec(x_44); -x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31; +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31; x_47 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_40, x_46, x_45); if (lean_obj_tag(x_47) == 0) { @@ -41834,7 +41802,7 @@ lean_object* x_48; lean_object* x_49; lean_object* x_50; x_48 = lean_ctor_get(x_47, 1); lean_inc(x_48); lean_dec(x_47); -x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32; +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32; x_50 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_40, x_49, x_48); if (lean_obj_tag(x_50) == 0) { @@ -41842,10 +41810,10 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33; +x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33; x_53 = l_Lean_Parser_Command_declVal___closed__2; -x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34; -x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35; +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34; +x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35; x_56 = l_Lean_Parser_registerAlias(x_52, x_53, x_54, x_55, x_6, x_51); if (lean_obj_tag(x_56) == 0) { @@ -41853,7 +41821,7 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_56, 1); lean_inc(x_57); lean_dec(x_56); -x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36; +x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_52, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -41861,7 +41829,7 @@ lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_59, 1); lean_inc(x_60); lean_dec(x_59); -x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_52, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -41869,10 +41837,10 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean x_63 = lean_ctor_get(x_62, 1); lean_inc(x_63); lean_dec(x_62); -x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38; +x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38; x_65 = l_Lean_Parser_Command_optDeclSig___closed__2; -x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39; -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40; +x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39; +x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40; x_68 = l_Lean_Parser_registerAlias(x_64, x_65, x_66, x_67, x_6, x_63); if (lean_obj_tag(x_68) == 0) { @@ -41880,7 +41848,7 @@ lean_object* x_69; lean_object* x_70; lean_object* x_71; x_69 = lean_ctor_get(x_68, 1); lean_inc(x_69); lean_dec(x_68); -x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41; +x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41; x_71 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_70, x_69); if (lean_obj_tag(x_71) == 0) { @@ -41888,7 +41856,7 @@ lean_object* x_72; lean_object* x_73; lean_object* x_74; x_72 = lean_ctor_get(x_71, 1); lean_inc(x_72); lean_dec(x_71); -x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42; x_74 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_73, x_72); if (lean_obj_tag(x_74) == 0) { @@ -41896,10 +41864,10 @@ lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean x_75 = lean_ctor_get(x_74, 1); lean_inc(x_75); lean_dec(x_74); -x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43; +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43; x_77 = l_Lean_Parser_Command_openDecl___closed__2; -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44; -x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45; +x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44; +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45; x_80 = l_Lean_Parser_registerAlias(x_76, x_77, x_78, x_79, x_6, x_75); if (lean_obj_tag(x_80) == 0) { @@ -41907,7 +41875,7 @@ lean_object* x_81; lean_object* x_82; lean_object* x_83; x_81 = lean_ctor_get(x_80, 1); lean_inc(x_81); lean_dec(x_80); -x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46; +x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46; x_83 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_76, x_82, x_81); if (lean_obj_tag(x_83) == 0) { @@ -41915,7 +41883,7 @@ lean_object* x_84; lean_object* x_85; lean_object* x_86; x_84 = lean_ctor_get(x_83, 1); lean_inc(x_84); lean_dec(x_83); -x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47; +x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47; x_86 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_76, x_85, x_84); if (lean_obj_tag(x_86) == 0) { @@ -41923,10 +41891,10 @@ lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean x_87 = lean_ctor_get(x_86, 1); lean_inc(x_87); lean_dec(x_86); -x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49; -x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50; -x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51; -x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52; +x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49; +x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50; +x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51; +x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52; x_92 = l_Lean_Parser_registerAlias(x_88, x_89, x_90, x_91, x_6, x_87); if (lean_obj_tag(x_92) == 0) { @@ -41934,7 +41902,7 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; x_93 = lean_ctor_get(x_92, 1); lean_inc(x_93); lean_dec(x_92); -x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53; +x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53; x_95 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_88, x_94, x_93); if (lean_obj_tag(x_95) == 0) { @@ -41942,7 +41910,7 @@ lean_object* x_96; lean_object* x_97; lean_object* x_98; x_96 = lean_ctor_get(x_95, 1); lean_inc(x_96); lean_dec(x_95); -x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54; +x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54; x_98 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_88, x_97, x_96); return x_98; } @@ -50905,8 +50873,6 @@ l_Lean_Parser_Command_include___closed__8 = _init_l_Lean_Parser_Command_include_ lean_mark_persistent(l_Lean_Parser_Command_include___closed__8); l_Lean_Parser_Command_include___closed__9 = _init_l_Lean_Parser_Command_include___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_include___closed__9); -l_Lean_Parser_Command_include___closed__10 = _init_l_Lean_Parser_Command_include___closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_include___closed__10); l_Lean_Parser_Command_include = _init_l_Lean_Parser_Command_include(); lean_mark_persistent(l_Lean_Parser_Command_include); if (builtin) {res = l___regBuiltin_Lean_Parser_Command_include__1(lean_io_mk_world()); @@ -50925,8 +50891,6 @@ l_Lean_Parser_Command_include_formatter___closed__3 = _init_l_Lean_Parser_Comman lean_mark_persistent(l_Lean_Parser_Command_include_formatter___closed__3); l_Lean_Parser_Command_include_formatter___closed__4 = _init_l_Lean_Parser_Command_include_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Command_include_formatter___closed__4); -l_Lean_Parser_Command_include_formatter___closed__5 = _init_l_Lean_Parser_Command_include_formatter___closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_include_formatter___closed__5); l___regBuiltin_Lean_Parser_Command_include_formatter__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_include_formatter__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_include_formatter__1___closed__1); l___regBuiltin_Lean_Parser_Command_include_formatter__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_include_formatter__1___closed__2(); @@ -50942,8 +50906,6 @@ l_Lean_Parser_Command_include_parenthesizer___closed__3 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_include_parenthesizer___closed__3); l_Lean_Parser_Command_include_parenthesizer___closed__4 = _init_l_Lean_Parser_Command_include_parenthesizer___closed__4(); lean_mark_persistent(l_Lean_Parser_Command_include_parenthesizer___closed__4); -l_Lean_Parser_Command_include_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_include_parenthesizer___closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_include_parenthesizer___closed__5); l___regBuiltin_Lean_Parser_Command_include_parenthesizer__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_include_parenthesizer__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_include_parenthesizer__1___closed__1); l___regBuiltin_Lean_Parser_Command_include_parenthesizer__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_include_parenthesizer__1___closed__2(); @@ -50995,122 +50957,122 @@ l_Lean_Parser_Command_eoi___closed__5 = _init_l_Lean_Parser_Command_eoi___closed lean_mark_persistent(l_Lean_Parser_Command_eoi___closed__5); l_Lean_Parser_Command_eoi = _init_l_Lean_Parser_Command_eoi(); lean_mark_persistent(l_Lean_Parser_Command_eoi); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3048_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3041_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Parser_Command_declModifiersF = _init_l_Lean_Parser_Command_declModifiersF(); lean_mark_persistent(l_Lean_Parser_Command_declModifiersF); l_Lean_Parser_Command_declModifiersT = _init_l_Lean_Parser_Command_declModifiersT(); lean_mark_persistent(l_Lean_Parser_Command_declModifiersT); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__50); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__51); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__52); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__53); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073____closed__54); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3073_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__50); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__51); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__52); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__53); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066____closed__54); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3066_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Parser_Term_open___closed__1 = _init_l_Lean_Parser_Term_open___closed__1(); diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index a413989e6386..bbf15bebfe0c 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -181,7 +181,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_RefInfo_toLspRefInfo LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_x27___at_Lean_Server_findReferences___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Server_References_allRefs___spec__7___boxed(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Server_combineIdents_buildIdMap___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Server_ModuleRefs_toLspModuleRefs___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_findReferences___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -11776,7 +11776,7 @@ x_4 = lean_apply_2(x_1, x_2, x_3); return x_4; } } -lean_object* l_Lean_Server_findModuleRefs(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_findModuleRefs(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; diff --git a/stage0/stdlib/Lean/Util/Path.c b/stage0/stdlib/Lean/Util/Path.c index aaae09a0e0cf..2b7ae1c95229 100644 --- a/stage0/stdlib/Lean/Util/Path.c +++ b/stage0/stdlib/Lean/Util/Path.c @@ -14,7 +14,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_List_foldl___at_Lean_moduleNameOfFileName___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SearchPath_findRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findSysroot___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_endsWith(lean_object*, lean_object*); @@ -28,19 +27,16 @@ lean_object* l_System_FilePath_walkDir(lean_object*, lean_object*, lean_object*) lean_object* l_System_FilePath_join(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_modToFilePath_go___spec__1(lean_object*); lean_object* l_System_FilePath_normalize(lean_object*); -static lean_object* l_Lean_moduleExists_go___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_forEachModuleInDir___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_searchModuleNameOfFileName___lambda__1(lean_object*, lean_object*, lean_object*); -lean_object* l_EStateM_instInhabited___rarg(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_SearchPath_findRoot(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_moduleNameOfFileName___lambda__2___closed__3; lean_object* l_System_SearchPath_parse(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_moduleNameOfFileName___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_System_FilePath_extension(lean_object*); +LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_modToFilePath_go___closed__3; static lean_object* l_Lean_findSysroot___lambda__1___closed__2; static lean_object* l_Lean_SearchPath_findAllWithExt___closed__1; @@ -53,7 +49,6 @@ static lean_object* l_Lean_findOLean___closed__4; LEAN_EXPORT lean_object* l_Lean_moduleNameOfFileName___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_FS_DirEntry_path(lean_object*); LEAN_EXPORT lean_object* l_Lean_forEachModuleInDir___rarg___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_findOLean___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* lean_module_name_of_file(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modToFilePath_go___boxed(lean_object*, lean_object*); @@ -65,37 +60,31 @@ lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__2___closed__1; static lean_object* l_Lean_findSysroot___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__1; -static lean_object* l_Lean_findOLean___closed__9; static lean_object* l_Lean_findOLean___closed__6; static lean_object* l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___closed__1; static lean_object* l_Lean_findOLean___closed__1; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initSearchPath(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -uint8_t l_String_isEmpty(lean_object*); lean_object* l_List_map___at_System_SearchPath_toString___spec__1(lean_object*); static lean_object* l_Lean_findOLean___closed__3; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findSysroot(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_forEachModuleInDir___rarg___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_System_FilePath_isDir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getLibDir___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_forEachModuleInDir___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getLibDir___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_findSysroot___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_moduleExists_go(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_findOLean___closed__7; static lean_object* l_Lean_modToFilePath_go___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___lambda__1(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_forEachModuleInDir___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_appDir(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_System_FilePath_components(lean_object*); @@ -105,14 +94,11 @@ LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt___boxed(lean_object static lean_object* l_Lean_getBuildDir___closed__2; lean_object* l_System_FilePath_addExtension(lean_object*, lean_object*); static lean_object* l_Lean_modToFilePath_go___closed__2; -extern lean_object* l_IO_instInhabitedError; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_initSearchPath___closed__1; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Path___hyg_1117_(lean_object*); static lean_object* l_Lean_getLibDir___closed__3; -LEAN_EXPORT lean_object* l_Lean_moduleExists___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Path___hyg_744_(lean_object*); LEAN_EXPORT lean_object* l_Lean_addSearchPathFromEnv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_searchModuleNameOfFileName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -127,9 +113,7 @@ static lean_object* l_Lean_getBuildDir___closed__4; static lean_object* l_Lean_findSysroot___lambda__1___closed__1; LEAN_EXPORT lean_object* lean_get_prefix(lean_object*); LEAN_EXPORT lean_object* l_Lean_SearchPath_findWithExt(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_moduleExists_go___closed__2; LEAN_EXPORT lean_object* l_Lean_searchModuleNameOfFileName(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SearchPath_findWithExt___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_internal_is_stage0(lean_object*); lean_object* lean_string_length(lean_object*); @@ -138,15 +122,11 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_searchPathRef; static lean_object* l_Lean_findOLean___closed__5; -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_findOLean___closed__8; -LEAN_EXPORT lean_object* l_Lean_findOLean(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_findOLean(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findSysroot___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_forEachModuleInDir___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_moduleExists_go___closed__1; extern lean_object* l_System_instInhabitedFilePath; static lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; uint8_t l_String_isPrefixOf(lean_object*, lean_object*); @@ -154,18 +134,13 @@ static lean_object* l_Lean_searchModuleNameOfFileName___closed__1; static lean_object* l_Lean_addSearchPathFromEnv___closed__1; lean_object* lean_io_realpath(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getBuiltinSearchPath(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_moduleExists_go___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modToFilePath_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_moduleExists(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_findOLean___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__3___boxed(lean_object*, lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); -lean_object* lean_io_read_dir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* lean_io_current_dir(lean_object*); static uint8_t l_Lean_getLibDir___closed__2; @@ -174,7 +149,6 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modToFilePath(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2(lean_object*, lean_object*, size_t, size_t); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_forEachModuleInDir(lean_object*); @@ -187,17 +161,14 @@ uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_getLibDir___lambda__1___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_moduleNameOfFileName___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* lean_get_libdir(lean_object*, lean_object*); static lean_object* l_Lean_SearchPath_findModuleWithExt___closed__1; lean_object* l_IO_Process_run(lean_object*, lean_object*); static lean_object* l_Lean_findSysroot___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_String_drop(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_moduleExists_go___spec__2___closed__1; static lean_object* l_Lean_modToFilePath_go___closed__4; LEAN_EXPORT lean_object* l_Lean_SearchPath_findAllWithExt(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_realPathNormalized(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_forEachModuleInDir___spec__1(lean_object* x_1, lean_object* x_2) { _start: @@ -809,712 +780,7 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_string_dec_eq(x_7, x_1); -lean_dec(x_7); -if (x_8 == 0) -{ -size_t x_9; size_t x_10; -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_3 = x_10; -goto _start; -} -else -{ -uint8_t x_12; -x_12 = 1; -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = 0; -return x_13; -} -} -} -static lean_object* _init_l_panic___at_Lean_moduleExists_go___spec__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_IO_instInhabitedError; -x_2 = lean_alloc_closure((void*)(l_EStateM_instInhabited___rarg), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_moduleExists_go___spec__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_panic___at_Lean_moduleExists_go___spec__2___closed__1; -x_4 = lean_panic_fn(x_3, x_1); -x_5 = lean_apply_1(x_4, x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -x_7 = l_Lean_modToFilePath(x_1, x_2, x_6); -x_8 = lean_io_read_dir(x_7, x_5); -lean_dec(x_7); -if (lean_obj_tag(x_8) == 0) -{ -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_array_get_size(x_10); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_nat_dec_lt(x_12, x_11); -if (x_13 == 0) -{ -uint8_t x_14; lean_object* x_15; -lean_dec(x_11); -lean_dec(x_10); -x_14 = 0; -x_15 = lean_box(x_14); -lean_ctor_set(x_8, 0, x_15); -return x_8; -} -else -{ -size_t x_16; size_t x_17; uint8_t x_18; lean_object* x_19; -x_16 = 0; -x_17 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_18 = l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1(x_3, x_10, x_16, x_17); -lean_dec(x_10); -x_19 = lean_box(x_18); -lean_ctor_set(x_8, 0, x_19); -return x_8; -} -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_20 = lean_ctor_get(x_8, 0); -x_21 = lean_ctor_get(x_8, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_8); -x_22 = lean_array_get_size(x_20); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_dec_lt(x_23, x_22); -if (x_24 == 0) -{ -uint8_t x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_22); -lean_dec(x_20); -x_25 = 0; -x_26 = lean_box(x_25); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_21); -return x_27; -} -else -{ -size_t x_28; size_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; -x_28 = 0; -x_29 = lean_usize_of_nat(x_22); -lean_dec(x_22); -x_30 = l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1(x_3, x_20, x_28, x_29); -lean_dec(x_20); -x_31 = lean_box(x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_21); -return x_32; -} -} -} -else -{ -uint8_t x_33; -x_33 = !lean_is_exclusive(x_8); -if (x_33 == 0) -{ -return x_8; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_8, 0); -x_35 = lean_ctor_get(x_8, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_8); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -} -} -static lean_object* _init_l_Lean_moduleExists_go___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(".", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_moduleExists_go___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.moduleExists.go", 20, 20); -return x_1; -} -} -static lean_object* _init_l_Lean_moduleExists_go___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_modToFilePath_go___closed__1; -x_2 = l_Lean_moduleExists_go___closed__2; -x_3 = lean_unsigned_to_nat(58u); -x_4 = lean_unsigned_to_nat(15u); -x_5 = l_Lean_modToFilePath_go___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_5; uint8_t x_6; -x_5 = l_System_FilePath_pathExists(x_1, x_4); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -return x_5; -} -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_5); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} -} -case 1: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_33; -x_10 = lean_ctor_get(x_3, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -lean_dec(x_3); -x_33 = l_String_isEmpty(x_2); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_34 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -x_35 = lean_string_append(x_34, x_11); -lean_dec(x_11); -x_36 = l_Lean_moduleExists_go___closed__1; -x_37 = lean_string_append(x_35, x_36); -x_38 = lean_string_append(x_37, x_2); -x_39 = lean_string_append(x_38, x_34); -x_12 = x_39; -goto block_32; -} -else -{ -x_12 = x_11; -goto block_32; -} -block_32: -{ -lean_object* x_13; lean_object* x_14; -x_13 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -lean_inc(x_10); -x_14 = l_Lean_moduleExists_go(x_1, x_13, x_10, x_4); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_unbox(x_15); -lean_dec(x_15); -if (x_16 == 0) -{ -uint8_t x_17; -lean_dec(x_12); -lean_dec(x_10); -x_17 = !lean_is_exclusive(x_14); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_14, 0); -lean_dec(x_18); -x_19 = 0; -x_20 = lean_box(x_19); -lean_ctor_set(x_14, 0, x_20); -return x_14; -} -else -{ -lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_dec(x_14); -x_22 = 0; -x_23 = lean_box(x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; -} -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_14, 1); -lean_inc(x_25); -lean_dec(x_14); -x_26 = lean_box(0); -x_27 = l_Lean_moduleExists_go___lambda__1(x_1, x_10, x_12, x_26, x_25); -lean_dec(x_12); -return x_27; -} -} -else -{ -uint8_t x_28; -lean_dec(x_12); -lean_dec(x_10); -x_28 = !lean_is_exclusive(x_14); -if (x_28 == 0) -{ -return x_14; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_14, 0); -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_14); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} -} -} -} -default: -{ -lean_object* x_40; lean_object* x_41; -lean_dec(x_3); -x_40 = l_Lean_moduleExists_go___closed__3; -x_41 = l_panic___at_Lean_moduleExists_go___spec__2(x_40, x_4); -return x_41; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_moduleExists_go___spec__1(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l_Lean_moduleExists_go___lambda__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_moduleExists_go(x_1, x_2, x_3, x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists(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_moduleExists_go(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_moduleExists___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_moduleExists(x_1, x_2, x_3, x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_string_dec_eq(x_7, x_1); -lean_dec(x_7); -if (x_8 == 0) -{ -size_t x_9; size_t x_10; -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_3 = x_10; -goto _start; -} -else -{ -uint8_t x_12; -x_12 = 1; -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = 0; -return x_13; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_string_dec_eq(x_7, x_1); -lean_dec(x_7); -if (x_8 == 0) -{ -size_t x_9; size_t x_10; -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_3 = x_10; -goto _start; -} -else -{ -uint8_t x_12; -x_12 = 1; -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = 0; -return x_13; -} -} -} -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -lean_inc(x_2); -lean_inc(x_1); -x_6 = l_System_FilePath_join(x_1, x_2); -x_7 = l_System_FilePath_isDir(x_6, x_5); -lean_dec(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -x_12 = lean_string_append(x_11, x_2); -lean_dec(x_2); -x_13 = l_Lean_moduleExists_go___closed__1; -x_14 = lean_string_append(x_12, x_13); -x_15 = lean_string_append(x_14, x_3); -x_16 = lean_string_append(x_15, x_11); -x_17 = lean_io_read_dir(x_1, x_10); -lean_dec(x_1); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_array_get_size(x_19); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_nat_dec_lt(x_21, x_20); -if (x_22 == 0) -{ -uint8_t x_23; lean_object* x_24; -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_16); -x_23 = 0; -x_24 = lean_box(x_23); -lean_ctor_set(x_17, 0, x_24); -return x_17; -} -else -{ -size_t x_25; size_t x_26; uint8_t x_27; lean_object* x_28; -x_25 = 0; -x_26 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_27 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1(x_16, x_19, x_25, x_26); -lean_dec(x_19); -lean_dec(x_16); -x_28 = lean_box(x_27); -lean_ctor_set(x_17, 0, x_28); -return x_17; -} -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_17, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_17); -x_31 = lean_array_get_size(x_29); -x_32 = lean_unsigned_to_nat(0u); -x_33 = lean_nat_dec_lt(x_32, x_31); -if (x_33 == 0) -{ -uint8_t x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_31); -lean_dec(x_29); -lean_dec(x_16); -x_34 = 0; -x_35 = lean_box(x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_30); -return x_36; -} -else -{ -size_t x_37; size_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; -x_37 = 0; -x_38 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_39 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1(x_16, x_29, x_37, x_38); -lean_dec(x_29); -lean_dec(x_16); -x_40 = lean_box(x_39); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_30); -return x_41; -} -} -} -else -{ -uint8_t x_42; -lean_dec(x_16); -x_42 = !lean_is_exclusive(x_17); -if (x_42 == 0) -{ -return x_17; -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_17, 0); -x_44 = lean_ctor_get(x_17, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_17); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; -} -} -} -else -{ -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_7, 1); -lean_inc(x_46); -lean_dec(x_7); -x_47 = lean_io_read_dir(x_1, x_46); -lean_dec(x_1); -if (lean_obj_tag(x_47) == 0) -{ -uint8_t x_48; -x_48 = !lean_is_exclusive(x_47); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_47, 0); -x_50 = lean_array_get_size(x_49); -x_51 = lean_unsigned_to_nat(0u); -x_52 = lean_nat_dec_lt(x_51, x_50); -if (x_52 == 0) -{ -uint8_t x_53; lean_object* x_54; -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_2); -x_53 = 0; -x_54 = lean_box(x_53); -lean_ctor_set(x_47, 0, x_54); -return x_47; -} -else -{ -size_t x_55; size_t x_56; uint8_t x_57; lean_object* x_58; -x_55 = 0; -x_56 = lean_usize_of_nat(x_50); -lean_dec(x_50); -x_57 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2(x_2, x_49, x_55, x_56); -lean_dec(x_49); -lean_dec(x_2); -x_58 = lean_box(x_57); -lean_ctor_set(x_47, 0, x_58); -return x_47; -} -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_47, 0); -x_60 = lean_ctor_get(x_47, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_47); -x_61 = lean_array_get_size(x_59); -x_62 = lean_unsigned_to_nat(0u); -x_63 = lean_nat_dec_lt(x_62, x_61); -if (x_63 == 0) -{ -uint8_t x_64; lean_object* x_65; lean_object* x_66; -lean_dec(x_61); -lean_dec(x_59); -lean_dec(x_2); -x_64 = 0; -x_65 = lean_box(x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_60); -return x_66; -} -else -{ -size_t x_67; size_t x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_67 = 0; -x_68 = lean_usize_of_nat(x_61); -lean_dec(x_61); -x_69 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2(x_2, x_59, x_67, x_68); -lean_dec(x_59); -lean_dec(x_2); -x_70 = lean_box(x_69); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_60); -return x_71; -} -} -} -else -{ -uint8_t x_72; -lean_dec(x_2); -x_72 = !lean_is_exclusive(x_47); -if (x_72 == 0) -{ -return x_47; -} -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_47, 0); -x_74 = lean_ctor_get(x_47, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_47); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_3) == 0) @@ -1529,193 +795,117 @@ return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); x_8 = lean_ctor_get(x_3, 1); lean_inc(x_8); lean_dec(x_3); -x_9 = l_System_FilePath_isDir(x_7, x_4); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_7); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_3 = x_8; -x_4 = x_12; -goto _start; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_dec(x_9); -x_15 = lean_box(0); lean_inc(x_2); lean_inc(x_7); -x_16 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1(x_7, x_2, x_1, x_15, x_14); -if (lean_obj_tag(x_16) == 0) +x_9 = l_System_FilePath_join(x_7, x_2); +x_10 = l_System_FilePath_isDir(x_9, x_4); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_System_FilePath_addExtension(x_9, x_1); +x_15 = l_System_FilePath_pathExists(x_14, x_13); +lean_dec(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_19; +lean_object* x_18; lean_dec(x_7); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); x_3 = x_8; -x_4 = x_19; +x_4 = x_18; goto _start; } else { -uint8_t x_21; +uint8_t x_20; lean_dec(x_8); lean_dec(x_2); -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_16, 0); -lean_dec(x_22); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_7); -lean_ctor_set(x_16, 0, x_23); -return x_16; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_15, 0); +lean_dec(x_21); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_7); +lean_ctor_set(x_15, 0, x_22); +return x_15; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_7); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 1); +lean_inc(x_23); +lean_dec(x_15); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_7); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } } else { -uint8_t x_27; +uint8_t x_26; +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) -{ -return x_16; +x_26 = !lean_is_exclusive(x_10); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_10, 0); +lean_dec(x_27); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_7); +lean_ctor_set(x_10, 0, x_28); +return x_10; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_10, 1); lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_SearchPath_findRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(x_2, x_3, x_1, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__1(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_SearchPath_findRoot___spec__2(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_6; -} +lean_dec(x_10); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_7); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; } -LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_SearchPath_findRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_SearchPath_findRoot(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; } } LEAN_EXPORT lean_object* l_Lean_SearchPath_findWithExt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; +lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_5 = l_Lean_Name_getRoot(x_3); x_6 = 0; x_7 = l_Lean_Name_toString(x_5, x_6); -x_8 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(x_2, x_7, x_1, x_4); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; +x_8 = l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1(x_2, x_7, x_1, x_4); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); if (lean_obj_tag(x_9) == 0) @@ -1808,29 +998,14 @@ return x_29; } } } -else -{ -uint8_t x_30; -lean_dec(x_3); -x_30 = !lean_is_exclusive(x_8); -if (x_30 == 0) -{ -return x_8; } -else +LEAN_EXPORT lean_object* l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_8, 0); -x_32 = lean_ctor_get(x_8, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_8); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} +lean_object* x_5; +x_5 = l_List_findM_x3f___at_Lean_SearchPath_findWithExt___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; } } LEAN_EXPORT lean_object* l_Lean_SearchPath_findWithExt___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -1856,227 +1031,125 @@ return x_4; static lean_object* _init_l_Lean_SearchPath_findModuleWithExt___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_SearchPath_findModuleWithExt___lambda__1___boxed), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; -x_5 = l_Lean_Name_getRoot(x_3); -x_6 = 0; -x_7 = l_Lean_Name_toString(x_5, x_6); -x_8 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(x_2, x_7, x_1, x_4); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_SearchPath_findModuleWithExt___closed__1; -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_3); -x_12 = lean_box(0); -x_13 = lean_apply_2(x_11, x_12, x_10); -return x_13; -} -else -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_9); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_9, 0); -lean_inc(x_3); -x_16 = l_Lean_moduleExists_go(x_15, x_2, x_3, x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_free_object(x_9); -lean_dec(x_15); -lean_dec(x_3); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = lean_box(0); -x_21 = lean_apply_2(x_11, x_20, x_19); -return x_21; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_SearchPath_findModuleWithExt___lambda__1___boxed), 2, 0); +return x_1; } -else +} +LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = l_Lean_SearchPath_findWithExt(x_1, x_2, x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l_Lean_SearchPath_findModuleWithExt___closed__1; +if (lean_obj_tag(x_6) == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_16, 0); -lean_dec(x_23); -x_24 = l_Lean_modToFilePath(x_15, x_3, x_2); -lean_dec(x_15); -lean_ctor_set(x_9, 0, x_24); -lean_ctor_set(x_16, 0, x_9); -return x_16; +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = lean_apply_2(x_8, x_9, x_7); +return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); -x_26 = l_Lean_modToFilePath(x_15, x_3, x_2); -lean_dec(x_15); -lean_ctor_set(x_9, 0, x_26); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_9); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} -} +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_6, 0); +x_13 = l_System_FilePath_pathExists(x_12, x_7); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_free_object(x_6); +lean_dec(x_12); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_box(0); +x_18 = lean_apply_2(x_8, x_17, x_16); +return x_18; } else { -uint8_t x_28; -lean_free_object(x_9); -lean_dec(x_15); -lean_dec(x_3); -x_28 = !lean_is_exclusive(x_16); -if (x_28 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_13); +if (x_19 == 0) { -return x_16; +lean_object* x_20; +x_20 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +lean_ctor_set(x_13, 0, x_6); +return x_13; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_16, 0); -x_30 = lean_ctor_get(x_16, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_16); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_9, 0); -lean_inc(x_32); -lean_dec(x_9); -lean_inc(x_3); -x_33 = l_Lean_moduleExists_go(x_32, x_2, x_3, x_10); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; uint8_t x_35; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_unbox(x_34); -lean_dec(x_34); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec(x_32); -lean_dec(x_3); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_box(0); -x_38 = lean_apply_2(x_11, x_37, x_36); -return x_38; -} -else +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_6, 0); +lean_inc(x_23); +lean_dec(x_6); +x_24 = l_System_FilePath_pathExists(x_23, x_7); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +if (x_26 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_39 = lean_ctor_get(x_33, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_33)) { - lean_ctor_release(x_33, 0); - lean_ctor_release(x_33, 1); - x_40 = x_33; -} else { - lean_dec_ref(x_33); - x_40 = lean_box(0); -} -x_41 = l_Lean_modToFilePath(x_32, x_3, x_2); -lean_dec(x_32); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -if (lean_is_scalar(x_40)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_40; -} -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_39); -return x_43; -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_23); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = lean_apply_2(x_8, x_28, x_27); +return x_29; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -lean_dec(x_32); -lean_dec(x_3); -x_44 = lean_ctor_get(x_33, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_33, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_33)) { - lean_ctor_release(x_33, 0); - lean_ctor_release(x_33, 1); - x_46 = x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_31 = x_24; } else { - lean_dec_ref(x_33); - x_46 = lean_box(0); + lean_dec_ref(x_24); + x_31 = lean_box(0); } -if (lean_is_scalar(x_46)) { - x_47 = lean_alloc_ctor(1, 2, 0); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_23); +if (lean_is_scalar(x_31)) { + x_33 = lean_alloc_ctor(0, 2, 0); } else { - x_47 = x_46; -} -lean_ctor_set(x_47, 0, x_44); -lean_ctor_set(x_47, 1, x_45); -return x_47; -} -} + x_33 = x_31; } +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +return x_33; } -else -{ -uint8_t x_48; -lean_dec(x_3); -x_48 = !lean_is_exclusive(x_8); -if (x_48 == 0) -{ -return x_8; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_8, 0); -x_50 = lean_ctor_get(x_8, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_8); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; } } } @@ -2374,7 +1447,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Path___hyg_1117_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Path___hyg_744_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -2918,445 +1991,128 @@ x_1 = lean_mk_string_unchecked("\n", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_findOLean___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("object file '", 13, 13); -return x_1; -} -} -static lean_object* _init_l_Lean_findOLean___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' of module ", 12, 12); -return x_1; -} -} -static lean_object* _init_l_Lean_findOLean___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" does not exist", 15, 15); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_findOLean(lean_object* x_1, uint8_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_findOLean(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_4 = l_Lean_initSearchPath___closed__1; -x_5 = lean_st_ref_get(x_4, x_3); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = l_Lean_Name_getRoot(x_1); -x_9 = 0; -x_10 = l_Lean_Name_toString(x_8, x_9); -x_11 = l_Lean_findOLean___closed__1; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = l_Lean_initSearchPath___closed__1; +x_4 = lean_st_ref_get(x_3, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); lean_inc(x_6); -lean_inc(x_10); -x_12 = l_List_findM_x3f___at_Lean_SearchPath_findRoot___spec__3(x_11, x_10, x_6, x_7); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = l_Lean_findOLean___closed__2; -x_17 = lean_string_append(x_16, x_10); -x_18 = l_Lean_findOLean___closed__3; -x_19 = lean_string_append(x_17, x_18); -x_20 = lean_string_append(x_19, x_10); -x_21 = l_Lean_findOLean___closed__4; -x_22 = lean_string_append(x_20, x_21); -x_23 = lean_string_append(x_22, x_10); -lean_dec(x_10); -x_24 = l_Lean_findOLean___closed__5; -x_25 = lean_string_append(x_23, x_24); -x_26 = l_List_map___at_System_SearchPath_toString___spec__1(x_6); -x_27 = l_Lean_findOLean___closed__6; -x_28 = l_String_intercalate(x_27, x_26); -x_29 = lean_string_append(x_25, x_28); -lean_dec(x_28); -x_30 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -x_31 = lean_string_append(x_29, x_30); -x_32 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 0, x_32); -return x_12; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_33 = lean_ctor_get(x_12, 1); -lean_inc(x_33); -lean_dec(x_12); -x_34 = l_Lean_findOLean___closed__2; -x_35 = lean_string_append(x_34, x_10); -x_36 = l_Lean_findOLean___closed__3; -x_37 = lean_string_append(x_35, x_36); -x_38 = lean_string_append(x_37, x_10); -x_39 = l_Lean_findOLean___closed__4; -x_40 = lean_string_append(x_38, x_39); -x_41 = lean_string_append(x_40, x_10); -lean_dec(x_10); -x_42 = l_Lean_findOLean___closed__5; -x_43 = lean_string_append(x_41, x_42); -x_44 = l_List_map___at_System_SearchPath_toString___spec__1(x_6); -x_45 = l_Lean_findOLean___closed__6; -x_46 = l_String_intercalate(x_45, x_44); -x_47 = lean_string_append(x_43, x_46); -lean_dec(x_46); -x_48 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; -x_49 = lean_string_append(x_47, x_48); -x_50 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_33); -return x_51; -} -} -else -{ -lean_object* x_52; uint8_t x_53; -lean_dec(x_10); -lean_dec(x_6); -x_52 = lean_ctor_get(x_12, 1); -lean_inc(x_52); -lean_dec(x_12); -x_53 = !lean_is_exclusive(x_13); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_13, 0); -lean_inc(x_1); -x_55 = l_Lean_modToFilePath(x_54, x_1, x_11); -lean_inc(x_1); -x_56 = l_Lean_moduleExists_go(x_54, x_11, x_1, x_52); -lean_dec(x_54); -if (lean_obj_tag(x_56) == 0) -{ -if (x_2 == 0) -{ -uint8_t x_57; -lean_free_object(x_13); -lean_dec(x_1); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) -{ -lean_object* x_58; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -lean_ctor_set(x_56, 0, x_55); -return x_56; -} -else -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_55); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} -} -else -{ -lean_object* x_61; uint8_t x_62; -x_61 = lean_ctor_get(x_56, 0); -lean_inc(x_61); -x_62 = lean_unbox(x_61); -lean_dec(x_61); -if (x_62 == 0) -{ -uint8_t x_63; -x_63 = !lean_is_exclusive(x_56); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_64 = lean_ctor_get(x_56, 0); -lean_dec(x_64); -x_65 = l_Lean_findOLean___closed__7; -x_66 = lean_string_append(x_65, x_55); -lean_dec(x_55); -x_67 = l_Lean_findOLean___closed__8; -x_68 = lean_string_append(x_66, x_67); -x_69 = 1; -x_70 = l_Lean_Name_toString(x_1, x_69); -x_71 = lean_string_append(x_68, x_70); -lean_dec(x_70); -x_72 = l_Lean_findOLean___closed__9; -x_73 = lean_string_append(x_71, x_72); -lean_ctor_set_tag(x_13, 18); -lean_ctor_set(x_13, 0, x_73); -lean_ctor_set_tag(x_56, 1); -lean_ctor_set(x_56, 0, x_13); -return x_56; -} -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_74 = lean_ctor_get(x_56, 1); -lean_inc(x_74); -lean_dec(x_56); -x_75 = l_Lean_findOLean___closed__7; -x_76 = lean_string_append(x_75, x_55); -lean_dec(x_55); -x_77 = l_Lean_findOLean___closed__8; -x_78 = lean_string_append(x_76, x_77); -x_79 = 1; -x_80 = l_Lean_Name_toString(x_1, x_79); -x_81 = lean_string_append(x_78, x_80); -lean_dec(x_80); -x_82 = l_Lean_findOLean___closed__9; -x_83 = lean_string_append(x_81, x_82); -lean_ctor_set_tag(x_13, 18); -lean_ctor_set(x_13, 0, x_83); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_13); -lean_ctor_set(x_84, 1, x_74); -return x_84; -} -} -else -{ -uint8_t x_85; -lean_free_object(x_13); -lean_dec(x_1); -x_85 = !lean_is_exclusive(x_56); -if (x_85 == 0) -{ -lean_object* x_86; -x_86 = lean_ctor_get(x_56, 0); -lean_dec(x_86); -lean_ctor_set(x_56, 0, x_55); -return x_56; -} -else -{ -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_56, 1); -lean_inc(x_87); -lean_dec(x_56); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_55); -lean_ctor_set(x_88, 1, x_87); -return x_88; -} -} -} -} -else -{ -uint8_t x_89; -lean_dec(x_55); -lean_free_object(x_13); -lean_dec(x_1); -x_89 = !lean_is_exclusive(x_56); -if (x_89 == 0) -{ -return x_56; -} -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_56, 0); -x_91 = lean_ctor_get(x_56, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_56); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; -} -} -} -else -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_13, 0); -lean_inc(x_93); -lean_dec(x_13); -lean_inc(x_1); -x_94 = l_Lean_modToFilePath(x_93, x_1, x_11); +lean_dec(x_4); +x_7 = l_Lean_findOLean___closed__1; lean_inc(x_1); -x_95 = l_Lean_moduleExists_go(x_93, x_11, x_1, x_52); -lean_dec(x_93); -if (lean_obj_tag(x_95) == 0) -{ -if (x_2 == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_1); -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_97 = x_95; -} else { - lean_dec_ref(x_95); - x_97 = lean_box(0); -} -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_97; -} -lean_ctor_set(x_98, 0, x_94); -lean_ctor_set(x_98, 1, x_96); -return x_98; -} -else +lean_inc(x_5); +x_8 = l_Lean_SearchPath_findWithExt(x_5, x_7, x_1, x_6); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_99; uint8_t x_100; -x_99 = lean_ctor_get(x_95, 0); -lean_inc(x_99); -x_100 = lean_unbox(x_99); -lean_dec(x_99); -if (x_100 == 0) -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_101 = lean_ctor_get(x_95, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_102 = x_95; -} else { - lean_dec_ref(x_95); - x_102 = lean_box(0); -} -x_103 = l_Lean_findOLean___closed__7; -x_104 = lean_string_append(x_103, x_94); -lean_dec(x_94); -x_105 = l_Lean_findOLean___closed__8; -x_106 = lean_string_append(x_104, x_105); -x_107 = 1; -x_108 = l_Lean_Name_toString(x_1, x_107); -x_109 = lean_string_append(x_106, x_108); -lean_dec(x_108); -x_110 = l_Lean_findOLean___closed__9; -x_111 = lean_string_append(x_109, x_110); -x_112 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_112, 0, x_111); -if (lean_is_scalar(x_102)) { - x_113 = lean_alloc_ctor(1, 2, 0); -} else { - x_113 = x_102; - lean_ctor_set_tag(x_113, 1); -} -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_101); -return x_113; -} -else +uint8_t x_10; +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = l_Lean_Name_getRoot(x_1); lean_dec(x_1); -x_114 = lean_ctor_get(x_95, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_115 = x_95; -} else { - lean_dec_ref(x_95); - x_115 = lean_box(0); -} -if (lean_is_scalar(x_115)) { - x_116 = lean_alloc_ctor(0, 2, 0); -} else { - x_116 = x_115; -} -lean_ctor_set(x_116, 0, x_94); -lean_ctor_set(x_116, 1, x_114); -return x_116; -} -} +x_13 = 0; +x_14 = l_Lean_Name_toString(x_12, x_13); +x_15 = l_Lean_findOLean___closed__2; +x_16 = lean_string_append(x_15, x_14); +x_17 = l_Lean_findOLean___closed__3; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_string_append(x_18, x_14); +x_20 = l_Lean_findOLean___closed__4; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_string_append(x_21, x_14); +lean_dec(x_14); +x_23 = l_Lean_findOLean___closed__5; +x_24 = lean_string_append(x_22, x_23); +x_25 = l_List_map___at_System_SearchPath_toString___spec__1(x_5); +x_26 = l_Lean_findOLean___closed__6; +x_27 = l_String_intercalate(x_26, x_25); +x_28 = lean_string_append(x_24, x_27); +lean_dec(x_27); +x_29 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_31); +return x_8; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_94); +lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_32 = lean_ctor_get(x_8, 1); +lean_inc(x_32); +lean_dec(x_8); +x_33 = l_Lean_Name_getRoot(x_1); lean_dec(x_1); -x_117 = lean_ctor_get(x_95, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_95, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_119 = x_95; -} else { - lean_dec_ref(x_95); - x_119 = lean_box(0); -} -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(1, 2, 0); -} else { - x_120 = x_119; -} -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_118); -return x_120; -} -} -} -} -else -{ -uint8_t x_121; -lean_dec(x_10); -lean_dec(x_6); +x_34 = 0; +x_35 = l_Lean_Name_toString(x_33, x_34); +x_36 = l_Lean_findOLean___closed__2; +x_37 = lean_string_append(x_36, x_35); +x_38 = l_Lean_findOLean___closed__3; +x_39 = lean_string_append(x_37, x_38); +x_40 = lean_string_append(x_39, x_35); +x_41 = l_Lean_findOLean___closed__4; +x_42 = lean_string_append(x_40, x_41); +x_43 = lean_string_append(x_42, x_35); +lean_dec(x_35); +x_44 = l_Lean_findOLean___closed__5; +x_45 = lean_string_append(x_43, x_44); +x_46 = l_List_map___at_System_SearchPath_toString___spec__1(x_5); +x_47 = l_Lean_findOLean___closed__6; +x_48 = l_String_intercalate(x_47, x_46); +x_49 = lean_string_append(x_45, x_48); +lean_dec(x_48); +x_50 = l_Array_forInUnsafe_loop___at_Lean_forEachModuleInDir___spec__2___rarg___lambda__5___closed__3; +x_51 = lean_string_append(x_49, x_50); +x_52 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_52, 0, x_51); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_32); +return x_53; +} +} +else +{ +uint8_t x_54; +lean_dec(x_5); lean_dec(x_1); -x_121 = !lean_is_exclusive(x_12); -if (x_121 == 0) +x_54 = !lean_is_exclusive(x_8); +if (x_54 == 0) { -return x_12; +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_8, 0); +lean_dec(x_55); +x_56 = lean_ctor_get(x_9, 0); +lean_inc(x_56); +lean_dec(x_9); +lean_ctor_set(x_8, 0, x_56); +return x_8; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_12, 0); -x_123 = lean_ctor_get(x_12, 1); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_12); -x_124 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -return x_124; -} -} +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_8, 1); +lean_inc(x_57); +lean_dec(x_8); +x_58 = lean_ctor_get(x_9, 0); +lean_inc(x_58); +lean_dec(x_9); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +return x_59; } } -LEAN_EXPORT lean_object* l_Lean_findOLean___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = lean_unbox(x_2); -lean_dec(x_2); -x_5 = l_Lean_findOLean(x_1, x_4, x_3); -return x_5; } } LEAN_EXPORT lean_object* l_List_foldl___at_Lean_moduleNameOfFileName___spec__1(lean_object* x_1, lean_object* x_2) { @@ -4100,21 +2856,13 @@ l_Lean_modToFilePath_go___closed__3 = _init_l_Lean_modToFilePath_go___closed__3( lean_mark_persistent(l_Lean_modToFilePath_go___closed__3); l_Lean_modToFilePath_go___closed__4 = _init_l_Lean_modToFilePath_go___closed__4(); lean_mark_persistent(l_Lean_modToFilePath_go___closed__4); -l_panic___at_Lean_moduleExists_go___spec__2___closed__1 = _init_l_panic___at_Lean_moduleExists_go___spec__2___closed__1(); -lean_mark_persistent(l_panic___at_Lean_moduleExists_go___spec__2___closed__1); -l_Lean_moduleExists_go___closed__1 = _init_l_Lean_moduleExists_go___closed__1(); -lean_mark_persistent(l_Lean_moduleExists_go___closed__1); -l_Lean_moduleExists_go___closed__2 = _init_l_Lean_moduleExists_go___closed__2(); -lean_mark_persistent(l_Lean_moduleExists_go___closed__2); -l_Lean_moduleExists_go___closed__3 = _init_l_Lean_moduleExists_go___closed__3(); -lean_mark_persistent(l_Lean_moduleExists_go___closed__3); l_Lean_SearchPath_findModuleWithExt___closed__1 = _init_l_Lean_SearchPath_findModuleWithExt___closed__1(); lean_mark_persistent(l_Lean_SearchPath_findModuleWithExt___closed__1); l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___closed__1 = _init_l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___closed__1(); lean_mark_persistent(l_List_forIn_loop___at_Lean_SearchPath_findAllWithExt___spec__2___closed__1); l_Lean_SearchPath_findAllWithExt___closed__1 = _init_l_Lean_SearchPath_findAllWithExt___closed__1(); lean_mark_persistent(l_Lean_SearchPath_findAllWithExt___closed__1); -if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Path___hyg_1117_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Path___hyg_744_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_searchPathRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_searchPathRef); @@ -4152,12 +2900,6 @@ l_Lean_findOLean___closed__5 = _init_l_Lean_findOLean___closed__5(); lean_mark_persistent(l_Lean_findOLean___closed__5); l_Lean_findOLean___closed__6 = _init_l_Lean_findOLean___closed__6(); lean_mark_persistent(l_Lean_findOLean___closed__6); -l_Lean_findOLean___closed__7 = _init_l_Lean_findOLean___closed__7(); -lean_mark_persistent(l_Lean_findOLean___closed__7); -l_Lean_findOLean___closed__8 = _init_l_Lean_findOLean___closed__8(); -lean_mark_persistent(l_Lean_findOLean___closed__8); -l_Lean_findOLean___closed__9 = _init_l_Lean_findOLean___closed__9(); -lean_mark_persistent(l_Lean_findOLean___closed__9); l_Lean_moduleNameOfFileName___lambda__2___closed__1 = _init_l_Lean_moduleNameOfFileName___lambda__2___closed__1(); lean_mark_persistent(l_Lean_moduleNameOfFileName___lambda__2___closed__1); l_Lean_moduleNameOfFileName___lambda__2___closed__2 = _init_l_Lean_moduleNameOfFileName___lambda__2___closed__2();